1 (**************************************************************************)
3 (* Landau's "Grundlagen der Analysis", formal specification for AUTOMATH *)
4 (* Copyright (C) 1977, L.S. van Benthem Jutting *)
5 (* 1992, revised by F. Wiedijk *)
6 (* 2008, 2014, revised by F. Guidi *)
8 (* Mechanized translation for COQ version 8 by F. Guidi *)
10 (**************************************************************************)
12 (* This file was generated by Helena 0.8.3 M (June 2015): do not edit *****)
15 Definition l_imp : (forall (a:Prop), (forall (b:Prop), Prop)).
16 exact (fun (a:Prop) => (fun (b:Prop) => (forall (x:a), b))).
20 Definition l_mp : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (i:l_imp a b), b)))).
21 exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:a) => (fun (i:l_imp a b) => i a1)))).
25 Definition l_refimp : (forall (a:Prop), l_imp a a).
26 exact (fun (a:Prop) => (fun (x:a) => x)).
30 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))))).
31 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))))))).
38 Definition l_not : (forall (a:Prop), Prop).
39 exact (fun (a:Prop) => l_imp a l_con).
43 Definition l_wel : (forall (a:Prop), Prop).
44 exact (fun (a:Prop) => l_not (l_not a)).
48 Definition l_weli : (forall (a:Prop), (forall (a1:a), l_wel a)).
49 exact (fun (a:Prop) => (fun (a1:a) => (fun (x:l_not a) => x a1))).
53 Axiom l_et : (forall (a:Prop), (forall (w:l_wel a), a)).
56 Definition l_cone : (forall (a:Prop), (forall (c1:l_con), a)).
57 exact (fun (a:Prop) => (fun (c1:l_con) => l_et a (fun (x:l_not a) => c1))).
61 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)))).
62 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))))))).
66 Definition l_imp_th2 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), l_imp a b))).
67 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)))).
71 Definition l_imp_th3 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not b), (forall (i:l_imp a b), l_not a)))).
72 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)))).
76 Definition l_imp_th4 : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (n:l_not b), l_not (l_imp a b))))).
77 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))))).
81 Definition l_imp_th5 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_imp a b)), a))).
82 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))))).
86 Definition l_imp_th6 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_imp a b)), l_not b))).
87 exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_imp a b)) => (fun (x:b) => n (fun (y:a) => x))))).
91 Definition l_imp_th7 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not b), (forall (i:l_imp (l_not a) b), a)))).
92 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))))).
96 Definition l_cp : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp (l_not b) (l_not a)), l_imp a b))).
97 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)))).
101 Definition l_obvious : Prop.
102 exact (l_imp l_con l_con).
106 Definition l_obviousi : l_obvious.
107 exact (l_refimp l_con).
111 Definition l_ec : (forall (a:Prop), (forall (b:Prop), Prop)).
112 exact (fun (a:Prop) => (fun (b:Prop) => l_imp a (l_not b))).
116 Definition l_eci1 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), l_ec a b))).
117 exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not a) => l_imp_th2 a (l_not b) n))).
121 Definition l_eci2 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not b), l_ec a b))).
122 exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not b) => (fun (x:a) => n)))).
126 Definition l_ec_th1 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp a (l_not b)), l_ec a b))).
127 exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp a (l_not b)) => i))).
131 Definition l_ec_th2 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp b (l_not a)), l_ec a b))).
132 exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp b (l_not a)) => (fun (x:a) => (fun (y:b) => i y x))))).
136 Definition l_comec : (forall (a:Prop), (forall (b:Prop), (forall (e:l_ec a b), l_ec b a))).
137 exact (fun (a:Prop) => (fun (b:Prop) => (fun (e:l_ec a b) => l_ec_th2 b a e))).
141 Definition l_ece1 : (forall (a:Prop), (forall (b:Prop), (forall (e:l_ec a b), (forall (a1:a), l_not b)))).
142 exact (fun (a:Prop) => (fun (b:Prop) => (fun (e:l_ec a b) => (fun (a1:a) => e a1)))).
146 Definition l_ece2 : (forall (a:Prop), (forall (b:Prop), (forall (e:l_ec a b), (forall (b1:b), l_not a)))).
147 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)))).
151 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))))).
152 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))))).
156 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))))).
157 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)))))).
161 Definition l_and : (forall (a:Prop), (forall (b:Prop), Prop)).
162 exact (fun (a:Prop) => (fun (b:Prop) => l_not (l_ec a b))).
166 Definition l_andi : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (b1:b), l_and a b)))).
167 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))))).
171 Definition l_ande1 : (forall (a:Prop), (forall (b:Prop), (forall (a1:l_and a b), a))).
172 exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:l_and a b) => l_imp_th5 a (l_not b) a1))).
176 Definition l_ande2 : (forall (a:Prop), (forall (b:Prop), (forall (a1:l_and a b), b))).
177 exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:l_and a b) => l_et b (l_imp_th6 a (l_not b) a1)))).
181 Definition l_comand : (forall (a:Prop), (forall (b:Prop), (forall (a1:l_and a b), l_and b a))).
182 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)))).
186 Definition l_and_th1 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), l_not (l_and a b)))).
187 exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not a) => l_weli (l_ec a b) (l_eci1 a b n)))).
191 Definition l_and_th2 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not b), l_not (l_and a b)))).
192 exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not b) => l_weli (l_ec a b) (l_eci2 a b n)))).
196 Definition l_and_th3 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_and a b)), (forall (a1:a), l_not b)))).
197 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)))).
201 Definition l_and_th4 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_and a b)), (forall (b1:b), l_not a)))).
202 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)))).
206 Definition l_and_th5 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_and a b)), l_not (l_and b a)))).
207 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)))).
211 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))))).
212 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)))))).
216 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))))).
217 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))))))).
221 Definition l_or : (forall (a:Prop), (forall (b:Prop), Prop)).
222 exact (fun (a:Prop) => (fun (b:Prop) => l_imp (l_not a) b)).
226 Definition l_ori1 : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), l_or a b))).
227 exact (fun (a:Prop) => (fun (b:Prop) => (fun (a1:a) => l_imp_th2 (l_not a) b (l_weli a a1)))).
231 Definition l_ori2 : (forall (a:Prop), (forall (b:Prop), (forall (b1:b), l_or a b))).
232 exact (fun (a:Prop) => (fun (b:Prop) => (fun (b1:b) => (fun (x:l_not a) => b1)))).
236 Definition l_or_th1 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp (l_not a) b), l_or a b))).
237 exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_imp (l_not a) b) => i))).
241 Definition l_or_th2 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp (l_not b) a), l_or a b))).
242 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))))).
246 Definition l_ore2 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), (forall (n:l_not a), b)))).
247 exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or a b) => (fun (n:l_not a) => o n)))).
251 Definition l_ore1 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), (forall (n:l_not b), a)))).
252 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))))).
256 Definition l_comor : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), l_or b a))).
257 exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or a b) => (fun (x:l_not b) => l_ore1 a b o x)))).
261 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))))).
262 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)))).
266 Definition l_or_th4 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_or a b)), l_not a))).
267 exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_or a b)) => l_imp_th5 (l_not a) b n))).
271 Definition l_or_th5 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not (l_or a b)), l_not b))).
272 exact (fun (a:Prop) => (fun (b:Prop) => (fun (n:l_not (l_or a b)) => l_imp_th6 (l_not a) b n))).
276 Definition l_or_th6 : (forall (a:Prop), l_or a (l_not a)).
277 exact (fun (a:Prop) => l_refimp (l_not a)).
281 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)))))).
282 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))))))).
286 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))))).
287 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))))).
291 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))))).
292 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))))).
296 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))))))).
297 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))))))).
301 Definition l_or_th10 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), l_imp (l_not a) b))).
302 exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or a b) => o))).
306 Definition l_or_th11 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or a b), l_imp (l_not b) a))).
307 exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_or a b) => l_comor a b o))).
311 Definition l_or_th12 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_or (l_not a) b), l_imp a b))).
312 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))).
316 Definition l_or_th13 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_imp a b), l_or (l_not a) b))).
317 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))).
321 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)))).
322 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)))).
326 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)))).
327 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)))).
331 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)))).
332 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)))).
336 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)))).
337 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)))).
341 Definition l_orec : (forall (a:Prop), (forall (b:Prop), Prop)).
342 exact (fun (a:Prop) => (fun (b:Prop) => l_and (l_or a b) (l_ec a b))).
346 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)))).
347 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)))).
351 Definition l_orec_th1 : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (n:l_not b), l_orec a b)))).
352 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))))).
356 Definition l_orec_th2 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), (forall (b1:b), l_orec a b)))).
357 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))))).
361 Definition l_orece1 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), l_or a b))).
362 exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => l_ande1 (l_or a b) (l_ec a b) o))).
366 Definition l_orece2 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), l_ec a b))).
367 exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => l_ande2 (l_or a b) (l_ec a b) o))).
371 Definition l_comorec : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), l_orec b a))).
372 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))))).
376 Definition l_orec_th3 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), (forall (a1:a), l_not b)))).
377 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)))).
381 Definition l_orec_th4 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), (forall (b1:b), l_not a)))).
382 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)))).
386 Definition l_orec_th5 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), (forall (n:l_not a), b)))).
387 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)))).
391 Definition l_orec_th6 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), (forall (n:l_not b), a)))).
392 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)))).
396 Definition l_iff : (forall (a:Prop), (forall (b:Prop), Prop)).
397 exact (fun (a:Prop) => (fun (b:Prop) => l_and (l_imp a b) (l_imp b a))).
401 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)))).
402 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)))).
406 Definition l_iff_th1 : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (b1:b), l_iff a b)))).
407 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))))).
411 Definition l_iff_th2 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), (forall (m:l_not b), l_iff a b)))).
412 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))))).
416 Definition l_iffe1 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_imp a b))).
417 exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => l_ande1 (l_imp a b) (l_imp b a) i))).
421 Definition l_iffe2 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_imp b a))).
422 exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => l_ande2 (l_imp a b) (l_imp b a) i))).
426 Definition l_comiff : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_iff b a))).
427 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)))).
431 Definition l_iff_th3 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), (forall (a1:a), b)))).
432 exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => (fun (a1:a) => l_iffe1 a b i a1)))).
436 Definition l_iff_th4 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), (forall (b1:b), a)))).
437 exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => (fun (b1:b) => l_iffe2 a b i b1)))).
441 Definition l_iff_th5 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), (forall (n:l_not a), l_not b)))).
442 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))))).
446 Definition l_iff_th6 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), (forall (n:l_not b), l_not a)))).
447 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))))).
451 Definition l_iff_th7 : (forall (a:Prop), (forall (b:Prop), (forall (a1:a), (forall (n:l_not b), l_not (l_iff a b))))).
452 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))))).
456 Definition l_iff_th8 : (forall (a:Prop), (forall (b:Prop), (forall (n:l_not a), (forall (b1:b), l_not (l_iff a b))))).
457 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))))).
461 Definition l_refiff : (forall (a:Prop), l_iff a a).
462 exact (fun (a:Prop) => l_iffi a a (l_refimp a) (l_refimp a)).
466 Definition l_symiff : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_iff b a))).
467 exact (fun (a:Prop) => (fun (b:Prop) => (fun (i:l_iff a b) => l_comiff a b i))).
471 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))))).
472 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))))))).
476 Definition l_iff_th9 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_imp (l_not a) (l_not b)))).
477 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)))).
481 Definition l_iff_th10 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_imp (l_not b) (l_not a)))).
482 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)))).
486 Definition l_iff_th11 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a b), l_iff (l_not a) (l_not b)))).
487 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)))).
491 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)))).
492 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))))).
496 Definition l_iff_th13 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), l_iff a (l_not b)))).
497 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))))).
501 Definition l_iff_th14 : (forall (a:Prop), (forall (b:Prop), (forall (o:l_orec a b), l_iff b (l_not a)))).
502 exact (fun (a:Prop) => (fun (b:Prop) => (fun (o:l_orec a b) => l_iff_th13 b a (l_comorec a b o)))).
506 Definition l_iff_th15 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff a (l_not b)), l_orec a b))).
507 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)))).
511 Definition l_iff_th16 : (forall (a:Prop), (forall (b:Prop), (forall (i:l_iff b (l_not a)), l_orec a b))).
512 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)))).
516 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))))).
517 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))))).
521 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))))).
522 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)))))).
526 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))))).
527 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)))))).
531 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))))).
532 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)))))).
536 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))))).
537 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)))))).
541 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))))).
542 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)))))).
546 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))))).
547 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)))))).
551 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))))).
552 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)))))).
556 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))))).
557 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))))))).
561 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))))).
562 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))))))).
566 Definition l_all : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Prop)).
567 exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (forall (x:sigma), p x))).
571 Definition l_alle : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (a1:l_all sigma p), (forall (s:sigma), p s)))).
572 exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (a1:l_all sigma p) => (fun (s:sigma) => a1 s)))).
576 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))))).
577 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)))))).
581 Definition l_non : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (x:sigma), Prop))).
582 exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (x:sigma) => l_not (p x)))).
586 Definition l_none : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Prop)).
587 exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => l_all sigma (l_non sigma p))).
591 Definition l_some : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Prop)).
592 exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => l_not (l_none sigma p))).
596 Definition l_somei : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (s:sigma), (forall (sp:p s), l_some sigma p)))).
597 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))))).
601 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))))).
602 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)))))).
606 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)))).
607 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))))).
611 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)))).
612 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)))).
616 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))))))).
617 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)))))).
621 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)))).
622 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))))).
626 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)))).
627 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)))).
631 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))).
632 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))).
636 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))))).
637 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)))).
641 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)))).
642 exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => (fun (n:l_none sigma p) => l_weli (l_none sigma p) n))).
646 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)))))))).
647 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)))))))).
651 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)))))).
652 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)))))))).
656 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))))).
657 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)))))).
661 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))))).
662 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))))))))).
666 Definition l_or3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), Prop))).
667 exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => l_or a (l_or b c)))).
671 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))))).
672 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))))).
676 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)))))).
677 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)))))).
681 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))))).
682 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)))))).
686 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)))))).
687 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)))))).
691 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))))).
692 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)))))).
696 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)))))).
697 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)))))).
701 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)))).
702 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))))).
706 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)))).
707 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))))).
711 Definition l_or3i1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:a), l_or3 a b c)))).
712 exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (a1:a) => l_ori1 a (l_or b c) a1)))).
716 Definition l_or3i2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (b1:b), l_or3 a b c)))).
717 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))))).
721 Definition l_or3i3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (c1:c), l_or3 a b c)))).
722 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))))).
726 Definition l_or3_th6 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or a b), l_or3 a b c)))).
727 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))))).
731 Definition l_or3_th7 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or b c), l_or3 a b c)))).
732 exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => (fun (o:l_or b c) => l_ori2 a (l_or b c) o)))).
736 Definition l_or3_th8 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_or c a), l_or3 a b c)))).
737 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))))).
741 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)))))))).
742 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))))))))).
746 Definition l_and3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), Prop))).
747 exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => l_and a (l_and b c)))).
751 Definition l_and3e1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), a)))).
752 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)))).
756 Definition l_and3e2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), b)))).
757 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))))).
761 Definition l_and3e3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), c)))).
762 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))))).
766 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)))))).
767 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))))))).
771 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)))).
772 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))))).
776 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)))).
777 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))))).
781 Definition l_and3_th3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), l_and a b)))).
782 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))))).
786 Definition l_and3_th4 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), l_and b c)))).
787 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)))).
791 Definition l_and3_th5 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (a1:l_and3 a b c), l_and c a)))).
792 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))))).
796 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)))).
797 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))))).
801 Definition l_ec3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), Prop))).
802 exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => l_and3 (l_ec a b) (l_ec b c) (l_ec c a)))).
806 Definition l_ec3_th1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), l_ec a b)))).
807 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)))).
811 Definition l_ec3_th2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), l_ec b c)))).
812 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)))).
816 Definition l_ec3_th3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (e:l_ec3 a b c), l_ec c a)))).
817 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)))).
821 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)))).
822 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)))).
826 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)))).
827 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))))).
831 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)))).
832 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)))))).
836 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))))).
837 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))))).
841 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))))).
842 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))))).
846 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))))).
847 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))))).
851 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))))).
852 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))))).
856 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))))).
857 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))))).
861 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))))).
862 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))))).
866 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)))))).
867 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)))))).
871 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))))).
872 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)))))).
876 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))))).
877 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))))).
881 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))))).
882 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))))).
886 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))))).
887 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)))))).
891 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))))).
892 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)))))).
896 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))))).
897 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)))))).
901 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)))))))))))).
902 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)))))))))))))).
906 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)))))))))))).
907 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)))))))))))))).
911 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)))))))))))).
912 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))))))))))))).
916 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)))))))))))).
917 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)))))))))))).
921 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)))))))))))).
922 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)))))))))))).
926 Definition l_orec3 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), Prop))).
927 exact (fun (a:Prop) => (fun (b:Prop) => (fun (c:Prop) => l_and (l_or3 a b c) (l_ec3 a b c)))).
931 Definition l_orec3e1 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_orec3 a b c), l_or3 a b c)))).
932 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)))).
936 Definition l_orec3e2 : (forall (a:Prop), (forall (b:Prop), (forall (c:Prop), (forall (o:l_orec3 a b c), l_ec3 a b c)))).
937 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)))).
941 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))))).
942 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))))).
946 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)))).
947 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)))))).
951 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)))).
952 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)))))).
956 Axiom l_e_is : (forall (sigma:Type), (forall (s:sigma), (forall (t:sigma), Prop))).
959 Axiom l_e_refis : (forall (sigma:Type), (forall (s:sigma), l_e_is sigma s s)).
962 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)))))).
965 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)))).
966 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)))).
970 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)))))).
971 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))))))).
975 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)))))).
976 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)))))).
980 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)))))).
981 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))))))).
985 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)))))).
986 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))))))).
990 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))))).
991 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))))).
995 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))))))).
996 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)))))).
1000 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))))))).
1001 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))))))).
1005 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))))))).
1006 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)))))).
1010 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))))))).
1011 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))))))).
1015 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))))))))).
1016 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)))))))).
1020 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)))))))).
1021 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)))))))).
1025 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)))))))))).
1026 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)))))))))).
1030 Definition l_e_amone : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Prop)).
1031 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)))))).
1035 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))))))).
1036 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))))))).
1040 Definition l_e_one : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Prop)).
1041 exact (fun (sigma:Type) => (fun (p:(forall (x:sigma), Prop)) => l_and (l_e_amone sigma p) (l_some sigma p))).
1045 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)))).
1046 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)))).
1050 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))).
1051 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))).
1055 Definition l_e_onee2 : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_one sigma p), l_some sigma p))).
1056 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))).
1060 Axiom l_e_ind : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_one sigma p), sigma))).
1063 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)))).
1066 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))))).
1067 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))))).
1071 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))))))).
1072 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)))))).
1076 Definition l_e_injective : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), Prop))).
1077 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)))))).
1081 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))))))).
1082 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))))))).
1086 Definition l_e_image : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (t0:tau), Prop)))).
1087 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)))))).
1091 Definition l_e_tofs : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), (forall (s:sigma), tau)))).
1092 exact (fun (sigma:Type) => (fun (tau:Type) => (fun (f:(forall (x:sigma), tau)) => (fun (s:sigma) => f s)))).
1096 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))))).
1097 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)))))).
1101 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)))))))))))).
1102 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))))))))))).
1106 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))))))))))).
1107 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)))))))))))).
1111 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))))))).
1112 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))))))))).
1116 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)))))))).
1117 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)))))).
1121 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)))))).
1122 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))))))).
1126 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)))))).
1127 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)))))).
1131 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)))))))).
1132 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))))))).
1136 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)))))).
1137 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))))))).
1141 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)))))))))).
1142 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)))))))))).
1146 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))))))))).
1147 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)))))))))).
1151 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))))))).
1152 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))))))).
1156 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))))).
1157 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)))))).
1161 Definition l_e_surjective : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), Prop))).
1162 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)))).
1166 Definition l_e_bijective : (forall (sigma:Type), (forall (tau:Type), (forall (f:(forall (x:sigma), tau)), Prop))).
1167 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)))).
1171 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)))).
1172 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)))).
1176 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)))).
1177 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)))).
1181 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))))).
1182 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)))))).
1186 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))))).
1187 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))))).
1191 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))))))).
1192 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)))))))).
1196 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))))))).
1197 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)))))).
1201 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)))))))).
1202 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))))))))).
1206 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))))))))))).
1207 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)))))))))).
1211 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)))))))))).
1212 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))))))))))).
1216 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))))))))).
1217 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)))))))))).
1221 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)))))))).
1222 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))))))))).
1226 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)))))))).
1227 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)))))))).
1231 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)))))))))).
1232 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)))))))))).
1236 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))))))))))))).
1237 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))))))))))))).
1241 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)))))))))))).
1242 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))))))))))))).
1246 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)))))))))).
1247 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)))))))))))).
1251 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)))))))).
1252 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)))))))))).
1256 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))))))))).
1257 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)))))))).
1261 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)))))))).
1262 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))))))))).
1266 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)))))))).
1267 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)))))))).
1271 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)))))))).
1272 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)))))))).
1276 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))))))))).
1277 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)))))))).
1281 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))))))).
1282 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)))))).
1286 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))))).
1289 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))))))))).
1290 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))))))))).
1294 Axiom l_e_ot : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), Type)).
1297 Axiom l_e_in : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), (forall (o1:l_e_ot sigma p), sigma))).
1300 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)))).
1303 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))).
1306 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)))).
1309 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)))))).
1310 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))))).
1314 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))))).
1315 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))))).
1319 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)))).
1320 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))))).
1324 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)))))))).
1325 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))))))).
1329 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))))))).
1330 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))))))).
1334 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))))).
1335 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)))))).
1339 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)))))).
1340 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))))).
1344 Axiom l_e_pairtype : (forall (sigma:Type), (forall (tau:Type), Type)).
1347 Axiom l_e_pair : (forall (sigma:Type), (forall (tau:Type), (forall (s:sigma), (forall (t:tau), l_e_pairtype sigma tau)))).
1350 Axiom l_e_first : (forall (sigma:Type), (forall (tau:Type), (forall (p1:l_e_pairtype sigma tau), sigma))).
1353 Axiom l_e_second : (forall (sigma:Type), (forall (tau:Type), (forall (p1:l_e_pairtype sigma tau), tau))).
1356 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))).
1359 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))))).
1360 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)))).
1364 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)))).
1367 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)))))).
1368 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))))).
1372 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)))).
1375 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)))))).
1376 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))))).
1380 Definition l_e_ite_prop1 : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), (forall (z:ksi), Prop))))).
1381 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))))))).
1385 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)))))))))).
1386 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))))))))).
1390 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))))))))).
1391 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)))))))))).
1395 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))))))))).
1396 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))))))))).
1400 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))))))))).
1401 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)))))))))).
1405 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)))))).
1406 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))))))))).
1410 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)))))).
1411 exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x:ksi) => (fun (y:ksi) => (fun (a1:a) => (fun (x1:a) => l_e_refis ksi x)))))).
1415 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)))))).
1416 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)))))).
1420 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))))).
1421 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)))))).
1425 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)))))).
1426 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)))))).
1430 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)))))).
1431 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)))))).
1435 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)))))))))).
1436 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))))))))).
1440 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))))))))).
1441 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)))))))))).
1445 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))))))))).
1446 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))))))))).
1450 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))))))))).
1451 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)))))))))).
1455 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)))))).
1456 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))))))))).
1460 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)))))).
1461 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)))))).
1465 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)))))).
1466 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))))).
1470 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))))).
1471 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)))))).
1475 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)))))).
1476 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)))))).
1480 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)))))).
1481 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)))))).
1485 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))))).
1486 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))))).
1490 Definition l_e_ite : (forall (a:Prop), (forall (ksi:Type), (forall (x:ksi), (forall (y:ksi), ksi)))).
1491 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))))).
1495 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))))).
1496 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))))).
1500 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))))).
1501 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))))).
1505 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))))).
1506 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))))).
1510 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))))).
1511 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)))))).
1515 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))))).
1516 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)))))).
1520 Definition l_e_wissel_wa : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), sigma)))).
1521 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)))).
1525 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))))).
1526 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))))).
1530 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))))).
1531 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))))).
1535 Definition l_e_wissel_wb : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (s:sigma), sigma)))).
1536 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))))).
1540 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))))).
1541 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))))).
1545 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)))))).
1546 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))))).
1550 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)))))).
1551 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)))))).
1555 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)))))).
1556 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))))))).
1560 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))))).
1561 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)))))).
1565 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)))))).
1566 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))))))).
1570 Definition l_e_wissel : (forall (sigma:Type), (forall (s0:sigma), (forall (t0:sigma), (forall (x:sigma), sigma)))).
1571 exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (x:sigma) => l_e_wissel_wb sigma s0 t0 x)))).
1575 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))))).
1576 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))))).
1580 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))))).
1581 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))))).
1585 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)))))).
1586 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)))))).
1590 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))))))))).
1591 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))))))))).
1595 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)))))))))).
1596 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))))))))).
1600 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))))))))).
1601 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))))))))))).
1605 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))))))))).
1606 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))))))))))).
1610 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))))))))).
1611 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))))))))).
1615 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))))))))).
1616 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)))))))))).
1620 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))))))))).
1621 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))))))))))).
1625 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))))))))).
1626 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))))))))).
1630 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)))))))).
1631 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)))))))))).
1635 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)))))))).
1636 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)))))))))).
1640 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)))))))).
1641 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)))))))).
1645 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)))))))).
1646 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)))))))).
1650 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))))))))).
1651 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))))))))).
1655 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)))))))).
1656 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)))))))))).
1660 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)))))))).
1661 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)))))))))).
1665 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)))))))).
1666 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)))))))).
1670 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)))))))).
1671 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)))))))).
1675 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))))))).
1676 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))))))))).
1680 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))))))).
1681 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))))))))).
1685 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)))))).
1686 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))))))).
1690 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)))).
1691 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)))))).
1695 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)))))).
1696 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))))))).
1700 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))))).
1701 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)))))).
1705 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)))))).
1706 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))))))).
1710 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))))).
1711 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)))))).
1715 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))))))).
1716 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))))))).
1720 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)))))).
1721 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))))))).
1725 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))))).
1726 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)))))).
1730 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)))).
1731 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))))).
1735 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)))).
1736 exact (fun (sigma:Type) => (fun (s0:sigma) => (fun (t0:sigma) => (fun (x:sigma) => l_e_wissel_t36 sigma s0 t0 x)))).
1740 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)))).
1741 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)))).
1745 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)))))).
1746 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))))))).
1750 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)))))))).
1751 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)))))))).
1755 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)))))))).
1756 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)))))))).
1760 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))))))))).
1761 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))))))))).
1765 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))))))).
1766 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)))))).
1770 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))))))).
1771 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)))))).
1775 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))))))).
1776 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)))))).
1780 Definition l_r_imp : (forall (a:Prop), (forall (b:(forall (x:a), Prop)), Prop)).
1781 exact (fun (a:Prop) => (fun (b:(forall (x:a), Prop)) => (forall (x:a), b x))).
1785 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)))).
1786 exact (fun (a:Prop) => (fun (b:(forall (x:a), Prop)) => (fun (a1:a) => (fun (i:l_r_imp a b) => i a1)))).
1790 Definition l_r_imp_th2 : (forall (a:Prop), (forall (b:(forall (x:a), Prop)), (forall (n:l_not a), l_r_imp a b))).
1791 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))))).
1795 Definition l_r_ec : (forall (a:Prop), (forall (b:(forall (x:a), Prop)), Prop)).
1796 exact (fun (a:Prop) => (fun (b:(forall (x:a), Prop)) => (forall (x:a), l_not (b x)))).
1800 Definition l_r_eci1 : (forall (a:Prop), (forall (b:(forall (x:a), Prop)), (forall (n:l_not a), l_r_ec a b))).
1801 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))))).
1805 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)))).
1806 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)))).
1810 Definition l_r_ite_is : (forall (a:Prop), (forall (ksi:Type), (forall (x1:ksi), (forall (y1:ksi), Prop)))).
1811 exact (fun (a:Prop) => (fun (ksi:Type) => (fun (x1:ksi) => (fun (y1:ksi) => l_e_is ksi x1 y1)))).
1815 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))))))).
1816 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)))))))))).
1820 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))))))))))))).
1821 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))))))))))).
1825 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)))))))))))).
1826 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)))))))))))).
1830 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)))))))))))).
1831 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))))))))))).
1835 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))))))))))).
1836 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)))))))))))).
1840 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)))))))).
1841 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))))))))))).
1845 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))))))))).
1846 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))))))).
1850 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))))))))).
1851 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)))))))).
1855 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)))))))).
1856 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)))))))).
1860 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)))))))).
1861 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)))))))).
1865 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)))))))).
1866 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)))))))).
1870 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))))))))))))).
1871 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))))))))))).
1875 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)))))))))))).
1876 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)))))))))))).
1880 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)))))))))))).
1881 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))))))))))).
1885 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))))))))))).
1886 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)))))))))))).
1890 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)))))))).
1891 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))))))))))).
1895 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))))))))).
1896 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))))))).
1900 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))))))))).
1901 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))))))).
1905 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)))))))).
1906 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)))))))).
1910 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)))))))).
1911 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)))))))).
1915 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)))))))).
1916 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)))))))).
1920 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))))))).
1921 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))))))).
1925 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)))))).
1926 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))))))).
1930 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))))))).
1931 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))))))).
1935 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)))))))).
1936 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))))))).
1940 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)))))))).
1941 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))))))).
1945 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)))))))).
1946 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)))))))).
1950 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)))))))).
1951 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)))))))).
1955 Axiom l_e_st_set : (forall (sigma:Type), Type).
1958 Axiom l_e_st_esti : (forall (sigma:Type), (forall (s:sigma), (forall (s0:l_e_st_set sigma), Prop))).
1961 Axiom l_e_st_setof : (forall (sigma:Type), (forall (p:(forall (x:sigma), Prop)), l_e_st_set sigma)).
1964 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))))).
1967 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)))).
1970 Definition l_e_st_empty : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), Prop)).
1971 exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => l_none sigma (fun (x:sigma) => l_e_st_esti sigma x s0))).
1975 Definition l_e_st_nonempty : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), Prop)).
1976 exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => l_some sigma (fun (x:sigma) => l_e_st_esti sigma x s0))).
1980 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))).
1981 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))).
1985 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))))).
1986 exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (e:l_e_st_empty sigma s0) => (fun (s:sigma) => e s)))).
1990 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)))).
1991 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)))).
1995 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))))).
1996 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))))).
2000 Definition l_e_st_incl : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), Prop))).
2001 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))))).
2005 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)))).
2006 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)))).
2010 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)))))).
2011 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)))))).
2015 Definition l_e_st_refincl : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), l_e_st_incl sigma s0 s0)).
2016 exact (fun (sigma:Type) => (fun (s0:l_e_st_set sigma) => (fun (x:sigma) => (fun (y:l_e_st_esti sigma x s0) => y)))).
2020 Definition l_e_st_disj : (forall (sigma:Type), (forall (s0:l_e_st_set sigma), (forall (t0:l_e_st_set sigma), Prop))).
2021 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))))).
2025 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)))).
2026 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)))).
2030 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)))).
2031 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)))))).
2035 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))))))).
2036 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)))))).
2040 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))))))).
2041 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)))))).
2045 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)))).
2046 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)))))).
2050 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))))))).
2051 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)))))))).
2055 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))))))).
2056 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)))))).
2060 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)))))).
2061 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)))))).
2065 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)))))).
2066 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)))))).
2070 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)))).
2071 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)))))).
2075 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)))).
2076 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)))))).
2080 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))))).
2083 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))))))).
2084 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))))))).
2088 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))))))).
2089 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))))))).
2093 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)))).
2094 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)))))).
2098 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)))))).
2099 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))))))).
2103 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)))))))).
2104 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))))))).
2108 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)))))))).
2109 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))))))).
2113 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)))))).
2114 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)))))))).
2118 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)))))).
2119 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)))))).
2123 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)))))).
2124 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))))).
2128 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))).
2129 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)))))).
2133 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))))))).
2134 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))))))).
2138 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))))))).
2139 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))))))).
2143 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)))))).
2144 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)))))).
2148 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)))))))).
2149 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)))))))).
2153 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)))))))))).
2154 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)))))))))).
2158 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)))))))))).
2159 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)))))))))).
2163 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)))))))))).
2164 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))))))))))).
2168 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)))))).
2169 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))))))).
2173 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))))))).
2174 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))))))).
2178 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))))))))).
2179 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)))))))).
2183 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)))))))).
2184 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)))))))).
2188 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))))))))))).
2189 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)))))))))))).
2193 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))))))))).
2194 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)))))))))).
2198 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)))))))))))).
2199 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)))))))))))).
2203 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))))))))).
2204 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)))))))))).
2208 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))))))).
2209 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))))))).
2213 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))))))).
2214 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)))))))).
2218 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)))))).
2219 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))))))).
2223 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))))))).
2224 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)))))))).
2228 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)))))))))))).
2229 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))))))))))).
2233 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)))))))))))).
2234 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))))))))))))).
2238 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)))))))))))).
2239 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)))))))))))).
2243 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)))))))))).
2244 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))))))))))).
2248 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))))))))))).
2249 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)))))))))))).
2253 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))))))))))).
2254 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)))))))))))).
2258 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))))))))).
2259 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)))))))))).
2263 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))))))).
2264 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))))))))).
2268 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)))))))))))))).
2269 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)))))))))))))))).
2273 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))))))))))))))).
2274 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))))))))))))))).
2278 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)))))))))))))).
2279 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))))))))))))))).
2283 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))))))))))).
2284 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))))))))))).
2288 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)))))))))))).
2289 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)))))))))))).
2293 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)))))))))).
2294 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))))))))))).
2298 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))))).
2299 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)))))).
2303 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))))))).
2304 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))))))).
2308 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)))))).
2309 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))))))).
2313 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)))))).
2314 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)))))).
2318 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))))))).
2319 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)))))).
2323 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))))))).
2324 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))))))).
2328 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)))))))).
2329 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)))))))).
2333 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)))))))).
2334 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))))))).
2338 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)))))))).
2339 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))))))).
2343 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)))))))).
2344 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))))))).
2348 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)))))))))).
2349 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)))))))))).
2353 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))))))))))).
2354 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)))))))))).
2358 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))))))))).
2359 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)))))))).
2363 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)))))))).
2364 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)))))))).
2368 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)))))))))))).
2369 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))))))))))))).
2373 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))))))))))).
2374 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)))))))))).
2378 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)))))))))))).
2379 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)))))))))))).
2383 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))))))))).
2384 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))))))))).
2388 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))))))).
2389 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))))))))))).
2393 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))))))))))).
2394 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)))))))))))).
2398 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)))))))))).
2399 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))))))))))).
2403 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))))))))))).
2404 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))))))))))))).
2408 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)))))))))))).
2409 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)))))))))))).
2413 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)))))))))))).
2414 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)))))))))))).
2418 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)))))))))).
2419 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))))))))))).
2423 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)))))))))))))))))).
2424 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))))))))))))))))).
2428 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)))))))))))))))))).
2429 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))))))))))))))))).
2433 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))))))))))))))))).
2434 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)))))))))))))))))).
2438 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))))))))))))))))).
2439 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))))))))))))))))).
2443 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))))))))))))))))).
2444 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))))))))))))))))).
2448 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))))))))))))))))).
2449 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)))))))))))))))))).
2453 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))))))))))))))).
2454 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))))))))))))))))).
2458 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))))))))))))).
2459 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))))))))))))))).
2463 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)))))))))).
2464 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))))))))))))).
2468 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)))))))))).
2469 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)))))))))).
2473 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))))))))).
2474 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)))))))))).
2478 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)))))))))))).
2479 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)))))))))).
2483 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)))))))))))).
2484 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)))))))))))).
2488 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))))))))))).
2489 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)))))))))).
2493 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))))))).
2494 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)))))))))))))).
2498 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))))))))))))).
2499 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))))))))))))).
2503 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)))))))))))).
2504 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)))))))))))).
2508 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)))))))))).
2509 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))))))))).
2513 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))))))))))))))).
2514 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)))))))))))))).
2518 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))))))))))))))).
2519 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)))))))))))))).
2523 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))))))))))))))).
2524 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)))))))))))))).
2528 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))))))))))))))).
2529 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)))))))))))))).
2533 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))))))))))))))).
2534 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))))))))))))))).
2538 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))))))))))))).
2539 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)))))))))))))).
2543 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)))))))))).
2544 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)))))))))).
2548 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))))))))))))))).
2549 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)))))))))))))).
2553 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))))))))))))))).
2554 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)))))))))))))).
2558 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))))))))))))))).
2559 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)))))))))))))).
2563 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))))))))))))))).
2564 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))))))))))))))).
2568 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)))))))))))).
2569 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))))))))))).
2573 Axiom l_e_st_eq_landau_n_nat : Type.
2576 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)).
2577 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)).
2581 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)).
2582 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))).
2586 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)).
2587 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)).
2591 Definition l_e_st_eq_landau_n_some : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), Prop).
2592 exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => l_some l_e_st_eq_landau_n_nat p).
2596 Definition l_e_st_eq_landau_n_all : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), Prop).
2597 exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => l_all l_e_st_eq_landau_n_nat p).
2601 Definition l_e_st_eq_landau_n_one : (forall (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)), Prop).
2602 exact (fun (p:(forall (x:l_e_st_eq_landau_n_nat), Prop)) => l_e_one l_e_st_eq_landau_n_nat p).
2606 Axiom l_e_st_eq_landau_n_1 : l_e_st_eq_landau_n_nat.
2609 Axiom l_e_st_eq_landau_n_suc : (forall (x:l_e_st_eq_landau_n_nat), l_e_st_eq_landau_n_nat).
2612 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)))).
2613 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))).
2617 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).
2620 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))).
2623 Definition l_e_st_eq_landau_n_cond1 : (forall (s:l_e_st_set l_e_st_eq_landau_n_nat), Prop).
2624 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).
2628 Definition l_e_st_eq_landau_n_cond2 : (forall (s:l_e_st_set l_e_st_eq_landau_n_nat), Prop).
2629 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))).
2633 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)))).
2636 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)))).
2637 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)))).
2641 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))))).
2642 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)))).
2646 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)))))).
2647 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)))))).
2651 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))))))).
2652 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)))))))).
2656 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))))).
2657 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)))).
2661 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)))).
2662 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))))).
2666 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)))).
2667 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)))).
2671 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)))).
2672 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)))).
2676 Definition l_e_st_eq_landau_n_22_prop1 : (forall (x:l_e_st_eq_landau_n_nat), Prop).
2677 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).
2681 Definition l_e_st_eq_landau_n_22_t1 : l_e_st_eq_landau_n_22_prop1 l_e_st_eq_landau_n_1.
2682 exact (l_e_st_eq_landau_n_ax3 l_e_st_eq_landau_n_1).
2686 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))).
2687 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)).
2691 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).
2692 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).
2696 Definition l_e_st_eq_landau_n_23_prop1 : (forall (x:l_e_st_eq_landau_n_nat), Prop).
2697 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)))).
2701 Definition l_e_st_eq_landau_n_23_t1 : l_e_st_eq_landau_n_23_prop1 l_e_st_eq_landau_n_1.
2702 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)).
2706 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))).
2707 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))).
2711 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)).
2712 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)).
2716 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).
2717 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).
2721 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)))).
2722 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)).
2726 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))))).
2727 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)))))).
2731 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))).
2732 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))))).
2736 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)))).
2737 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))).
2741 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)).
2742 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))))).
2746 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)).
2747 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))).
2751 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)))))).
2752 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))))))).
2756 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)))))).
2757 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))))).
2761 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)))))).
2762 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))))).
2766 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))))).
2767 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)))))).
2771 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))))))))).
2772 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))))))).
2776 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))))))).
2777 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))))))).
2781 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))))))).
2782 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))))))).
2786 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))))))))).
2787 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))))))).
2791 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))))))))).
2792 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))))))).
2796 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)))))))).
2797 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))))))))).
2801 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)))))).
2802 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)))))).
2806 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))))).
2807 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)))))).
2811 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)).
2812 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))))).
2816 Definition l_e_st_eq_landau_n_24_prop4 : (forall (x:l_e_st_eq_landau_n_nat), Prop).
2817 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)).
2821 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.
2822 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))).
2826 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.
2827 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).
2831 Definition l_e_st_eq_landau_n_24_t14 : l_e_st_eq_landau_n_24_prop4 l_e_st_eq_landau_n_1.
2832 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).
2836 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))))).
2837 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)))))).
2841 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))))))).
2842 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)))))).
2846 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))))).
2847 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)))).
2851 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)))))).
2852 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)))))).
2856 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))))).
2857 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))))).
2861 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))))))).
2862 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))))).
2866 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)))))).
2867 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)))))).
2871 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))))))).
2872 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))))))).
2876 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))))).
2877 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))))).
2881 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))))).
2882 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))))).
2886 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))))).
2887 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))))).
2891 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))).
2892 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)))).
2896 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).
2897 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).
2901 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)))))).
2902 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)).
2906 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)).
2907 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)).
2911 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)).
2912 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)).
2916 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)).
2917 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)).
2921 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)).
2922 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)).
2926 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)).
2927 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)).
2931 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)))).
2932 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)).
2936 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.
2937 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).
2941 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)).
2942 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).
2946 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))).
2947 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))).
2951 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)))).
2952 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)).
2956 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)).
2957 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)).
2961 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)))).
2962 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))).
2966 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)).
2967 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)).
2971 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))).
2972 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))).
2976 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))))).
2977 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)))).
2981 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))))).
2982 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)))).
2986 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))))))).
2987 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))))))).
2991 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))).
2992 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))))).
2996 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)).
2997 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)))).
3001 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))))))).
3002 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)))).
3006 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))))).
3007 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)))))).
3011 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))))).
3012 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))).
3016 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))))).
3017 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))).
3021 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)))).
3022 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)))).
3026 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)).
3027 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))).
3031 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))).
3032 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)).
3036 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))).
3037 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)).
3041 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)).
3042 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))).
3046 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))))).
3047 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)))).
3051 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))))).
3052 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))).
3056 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))).
3057 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)))).
3061 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))).
3062 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)).
3066 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))).
3067 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)).
3071 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).
3072 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)).
3076 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)))).
3077 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)))).
3081 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))).
3082 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)).
3086 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)).
3087 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))).
3091 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)).
3092 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)).
3096 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).
3097 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)).
3101 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))))).
3102 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))).
3106 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)))).
3107 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)))).
3111 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))).
3112 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)).
3116 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)))).
3117 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))))).
3121 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))))).
3122 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)))).
3126 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)))).
3127 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))))).
3131 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))))))).
3132 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))))).
3136 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))))).
3137 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)))))).
3141 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))))).
3142 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)))).
3146 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)))).
3147 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))))).
3151 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))).
3152 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)))).
3156 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)))))).
3157 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))))))).
3161 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)))).
3162 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)))))).
3166 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)).
3167 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)).
3171 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)).
3172 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))).
3176 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)).
3177 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))).
3181 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))))).
3182 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))))).
3186 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))))).
3187 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))))).
3191 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)))).
3192 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)))).
3196 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))).
3197 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))).
3201 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)))).
3202 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)))).
3206 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))).
3207 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))).
3211 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))))))))).
3212 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)))))))))).
3216 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)))))))).
3217 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))))))))).
3221 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)))))).
3222 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)))))))).
3226 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)))).
3227 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)))))).
3231 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)))).
3232 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)))).
3236 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))).
3237 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))).
3241 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))).
3242 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))).
3246 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)).
3247 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))).
3251 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)).
3252 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)).
3256 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))))).
3257 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))))).
3261 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)))).
3262 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))))).
3266 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)))).
3267 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)))))).
3271 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)))).
3272 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))))).
3276 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)).
3277 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)))).
3281 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).
3282 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)).
3286 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))))).
3287 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)))))).
3291 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))))).
3292 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))))).
3296 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))))).
3297 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))))).
3301 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)))))))).
3302 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)))))))).
3306 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)))))))).
3307 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)))))))).
3311 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)))))))))).
3312 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)))))))))).
3316 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)))))))))).
3317 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))))))))))).
3321 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)))))))))).
3322 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)))))))))).
3326 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)))))))).
3327 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))))))))).
3331 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)))))))).
3332 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)))))))).
3336 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))))))).
3337 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))))))).
3341 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))))).
3342 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)))))).
3346 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)))))))).
3347 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))))))).
3351 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))))))).
3352 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))))))).
3356 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))))).
3357 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)))))).
3361 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))))).
3362 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))))).
3366 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)))).
3367 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)))).
3371 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))).
3372 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)).
3376 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))))).
3377 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))).
3381 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)))).
3382 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)).
3386 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)))).
3387 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)).
3391 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)).
3392 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))).
3396 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)).
3397 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))).
3401 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))).
3402 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)).
3406 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))).
3407 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)).
3411 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))).
3412 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)).
3416 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))).
3417 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))).
3421 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))).
3422 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))).
3426 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)).
3427 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))).
3431 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)).
3432 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))).
3436 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))).
3437 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)))).
3441 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))).
3442 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)))).
3446 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))))).
3447 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))))).
3451 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))))).
3452 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))))).
3456 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))))).
3457 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))))).
3461 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))))).
3462 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))))).
3466 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))))).
3467 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))))).
3471 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))))).
3472 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))))).
3476 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))))).
3477 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))))).
3481 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))))).
3482 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))))).
3486 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))).
3487 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))).
3491 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))).
3492 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))).
3496 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))).
3497 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))).
3501 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))).
3502 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))).
3506 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))))))).
3507 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)))))))).
3511 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))))))).
3512 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)))))))).
3516 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))))))).
3517 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)))))))).
3521 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))))))).
3522 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)))))))).
3526 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)))).
3527 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)))).
3531 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)))).
3532 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))).
3536 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))).
3537 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))).
3541 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))).
3542 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)))).
3546 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)))).
3547 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)))).
3551 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)))).
3552 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)))).
3556 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))).
3557 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)))).
3561 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))).
3562 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)))).
3566 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))))))))))).
3567 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)))))))))).
3571 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))))))))).
3572 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)))))))))).
3576 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))))))).
3577 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))))))))).
3581 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))))).
3582 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))))))).
3586 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))))).
3587 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))))).
3591 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))))).
3592 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))))).
3596 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))))).
3597 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))))))).
3601 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))))).
3602 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)))))).
3606 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))))).
3607 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)))))).
3611 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))))).
3612 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)))))).
3616 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))))).
3617 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))))).
3621 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))))))).
3622 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)))))))).
3626 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))))))).
3627 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)))))))).
3631 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)))))).
3632 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))))))).
3636 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)))))).
3637 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))))))).
3641 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))))).
3642 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)))))).
3646 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)))))).
3647 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))))))).
3651 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)))))).
3652 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)))))).
3656 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))))).
3657 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)))))).
3661 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))))).
3662 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))))).
3666 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))))).
3667 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))))))).
3671 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)).
3672 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)))).
3676 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))).
3677 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)).
3681 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).
3682 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)).
3686 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)).
3687 exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz18b x).
3691 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))))))).
3692 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))))))).
3696 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))))))).
3697 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)))))))).
3701 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))))))).
3702 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))))))).
3706 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))))).
3707 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)))))).
3711 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))))).
3712 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)))).
3716 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))))).
3717 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)))))).
3721 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))))).
3722 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)))).
3726 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))))).
3727 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))))).
3731 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))))).
3732 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)))).
3736 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))))).
3737 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))))).
3741 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))))).
3742 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)))).
3746 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))))))).
3747 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))))))).
3751 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))))))).
3752 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))))))).
3756 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))))))).
3757 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))))))).
3761 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))))))).
3762 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))))))).
3766 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)))))).
3767 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)))))).
3771 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)))))).
3772 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)))))).
3776 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))))).
3777 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))))).
3781 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))))).
3782 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))))).
3786 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))))).
3787 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)))))).
3791 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))))).
3792 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)))))).
3796 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)))).
3797 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))).
3801 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))))).
3802 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)))).
3806 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)))).
3807 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)))).
3811 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)))).
3812 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)))).
3816 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)))).
3817 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)))).
3821 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))))).
3822 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))))).
3826 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)))).
3827 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))))).
3831 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)))).
3832 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)))).
3836 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)))).
3837 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))))).
3841 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)))).
3842 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)))))).
3846 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)))).
3847 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))))).
3851 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))))))).
3852 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)))))).
3856 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))))))).
3857 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))))))).
3861 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))))))).
3862 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))))))).
3866 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))))))).
3867 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))))))).
3871 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))))))).
3872 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)))))).
3876 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))))))).
3877 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)))))))).
3881 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))))))).
3882 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))))))).
3886 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))))))).
3887 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))))))).
3891 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))))))).
3892 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)))))).
3896 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))))))).
3897 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))))))).
3901 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))))))))).
3902 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)))))))))).
3906 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))))))))).
3907 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))))))))).
3911 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)))))))).
3912 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)))))))).
3916 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)))))))).
3917 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)))))))).
3921 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))))))).
3922 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))))))).
3926 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)))))))).
3927 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)))))))).
3931 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)))))))).
3932 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)))))))).
3936 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))))))).
3937 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))))))).
3941 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))))))).
3942 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)))))))).
3946 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))))).
3947 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))))).
3951 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)))).
3952 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))))).
3956 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)).
3957 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)))).
3961 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).
3962 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)).
3966 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).
3967 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)).
3971 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).
3972 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)).
3976 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)).
3977 exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz24b x).
3981 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)))))).
3982 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)))))).
3986 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)))))).
3987 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)))))).
3991 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)))).
3992 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))))).
3996 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)))).
3997 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)))).
4001 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))).
4002 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)))).
4006 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))).
4007 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)))).
4011 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))))).
4012 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)))).
4016 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)))).
4017 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)))).
4021 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))).
4022 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)))).
4026 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))).
4027 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)))).
4031 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))).
4032 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)))).
4036 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))).
4037 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)))).
4041 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))).
4042 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)))).
4046 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)).
4047 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))).
4051 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)).
4052 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))).
4056 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))).
4057 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)))).
4061 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)).
4062 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))).
4066 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))))).
4067 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))))).
4071 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)))))).
4072 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)))))).
4076 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)))))).
4077 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)))))).
4081 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))))))).
4082 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)))))).
4086 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))))).
4087 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)))))).
4091 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))).
4092 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))))).
4096 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))))))))).
4097 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))))).
4101 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)))))).
4102 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)))))).
4106 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)))))).
4107 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)))))).
4111 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)))).
4112 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)))).
4116 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)))))).
4117 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)))).
4121 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)))).
4122 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)))).
4126 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)))))).
4127 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)))).
4131 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))))))).
4132 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)))))))).
4136 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)))))))).
4137 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))))))))).
4141 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))))))).
4142 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)))))))).
4146 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))))))).
4147 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)))))))).
4151 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)))))).
4152 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))))))).
4156 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))))).
4157 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)))))).
4161 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)))).
4162 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))))).
4166 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)))).
4167 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))))).
4171 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))).
4172 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)))).
4176 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))).
4177 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)))).
4181 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)).
4182 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))).
4186 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))))).
4187 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)))).
4191 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))))).
4192 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)))).
4196 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)))))).
4197 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))))))).
4201 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)))))).
4202 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))))))).
4206 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)))))).
4207 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))))))).
4211 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)))))).
4212 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))))))).
4216 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))))).
4217 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)))))).
4221 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))))).
4222 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))))).
4226 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))).
4227 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))).
4231 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))))).
4232 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))))).
4236 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))))).
4237 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))))).
4241 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)))))).
4242 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))))).
4246 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))))).
4247 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)))))).
4251 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))).
4252 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)))).
4256 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))))).
4257 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))))).
4261 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))))).
4262 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))))).
4266 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))))).
4267 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))))).
4271 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))))).
4272 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))))).
4276 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))))).
4277 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))))).
4281 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))))).
4282 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))))).
4286 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))))).
4287 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)))))).
4291 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))))).
4292 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)))))).
4296 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))))).
4297 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))))))).
4301 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)).
4302 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))))).
4306 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))).
4307 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))).
4311 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)).
4312 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)))).
4316 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)).
4317 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))).
4321 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)))))).
4322 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))))))).
4326 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))))).
4327 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))))).
4331 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))))).
4332 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))))).
4336 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))))).
4337 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)))))).
4341 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)))))))).
4342 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))))))).
4346 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))))))).
4347 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))))))).
4351 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))))))).
4352 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))))))).
4356 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)))))))).
4357 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))))))).
4361 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)))))))).
4362 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))))))).
4366 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)))))))).
4367 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))))))))).
4371 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)))))).
4372 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)))))).
4376 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))))).
4377 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)))))).
4381 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)).
4382 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))))).
4386 Definition l_e_st_eq_landau_n_428_prop4 : (forall (x:l_e_st_eq_landau_n_nat), Prop).
4387 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)).
4391 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).
4392 exact (fun (y:l_e_st_eq_landau_n_nat) => y).
4396 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.
4397 exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_satz4e x).
4401 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.
4402 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).
4406 Definition l_e_st_eq_landau_n_428_t14 : l_e_st_eq_landau_n_428_prop4 l_e_st_eq_landau_n_1.
4407 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).
4411 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))))).
4412 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))))).
4416 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)))).
4417 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)))).
4421 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))))).
4422 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))))).
4426 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))))).
4427 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))))).
4431 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)))))).
4432 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))))).
4436 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)))))))).
4437 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))))))).
4441 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))))))).
4442 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)))))).
4446 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))))))).
4447 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))))))).
4451 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))))).
4452 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))))).
4456 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))))).
4457 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))))).
4461 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))))).
4462 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))))).
4466 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))).
4467 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)))).
4471 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).
4472 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).
4476 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))))).
4477 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)).
4481 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)).
4482 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)).
4486 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)).
4487 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)).
4491 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)).
4492 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)).
4496 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).
4497 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)).
4501 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)).
4502 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)).
4506 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))).
4507 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)).
4511 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.
4512 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).
4516 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).
4517 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).
4521 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)).
4522 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))).
4526 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))).
4527 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)).
4531 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)).
4532 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)).
4536 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)))).
4537 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))).
4541 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)).
4542 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)).
4546 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))).
4547 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))).
4551 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))))).
4552 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)))).
4556 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))))).
4557 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)))).
4561 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))))))).
4562 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))))))).
4566 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)).
4567 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))).
4571 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)).
4572 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)).
4576 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)).
4577 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)).
4581 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)).
4582 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))).
4586 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))))).
4587 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)))).
4591 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)))).
4592 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))).
4596 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))).
4597 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)))).
4601 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))).
4602 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)).
4606 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))).
4607 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)).
4611 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).
4612 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)).
4616 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)))).
4617 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)))).
4621 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))).
4622 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)).
4626 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))).
4627 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))))).
4631 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)).
4632 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)))).
4636 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))))).
4637 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))))).
4641 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))))).
4642 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)))))).
4646 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))))).
4647 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))).
4651 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))))).
4652 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))))).
4656 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))))).
4657 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))).
4661 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)))).
4662 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)))).
4666 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))))).
4667 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)))).
4671 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))).
4672 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))))).
4676 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)).
4677 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)))).
4681 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))))).
4682 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)))))).
4686 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))))).
4687 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))).
4691 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))))).
4692 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))).
4696 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)))).
4697 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)))).
4701 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)))))))).
4702 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))))))).
4706 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))))))).
4707 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))))))).
4711 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))))).
4712 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)))))).
4716 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))))).
4717 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)))).
4721 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))))).
4722 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)))))).
4726 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))))).
4727 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)))).
4731 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))))).
4732 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))))).
4736 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))))).
4737 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)))).
4741 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))))).
4742 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))))).
4746 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))))).
4747 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)))).
4751 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))))))).
4752 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))))))).
4756 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))))))).
4757 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))))))).
4761 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))))))).
4762 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))))))).
4766 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))))))).
4767 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))))))).
4771 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)))))).
4772 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)))))).
4776 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)))))).
4777 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)))))).
4781 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))))).
4782 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))))).
4786 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))))).
4787 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))))).
4791 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))))).
4792 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)))))).
4796 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))))).
4797 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)))))).
4801 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)))).
4802 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))).
4806 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))))).
4807 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)))).
4811 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)))).
4812 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)))).
4816 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)))).
4817 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)))).
4821 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)))).
4822 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)))).
4826 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)))).
4827 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)))).
4831 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))))))).
4832 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)))))).
4836 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))))))).
4837 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))))))).
4841 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))))))).
4842 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))))))).
4846 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))))))).
4847 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))))))).
4851 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))))))).
4852 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)))))).
4856 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))))))).
4857 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)))))))).
4861 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))))))).
4862 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))))))).
4866 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))))))).
4867 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))))))).
4871 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))))))).
4872 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)))))).
4876 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))))))).
4877 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))))))).
4881 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))))))))).
4882 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)))))))))).
4886 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))))))))).
4887 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))))))))).
4891 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)))))))).
4892 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)))))))).
4896 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)))))))).
4897 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)))))))).
4901 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))))))).
4902 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))))))).
4906 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)))))))).
4907 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)))))))).
4911 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)))))))).
4912 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)))))))).
4916 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))))))).
4917 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))))))).
4921 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))))))).
4922 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)))))))).
4926 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)))).
4927 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))).
4931 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))).
4932 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)))).
4936 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))))).
4937 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)))).
4941 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))).
4942 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)))).
4946 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)))).
4947 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))))).
4951 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))).
4952 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)))).
4956 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)))))).
4957 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)))))).
4961 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)))))))).
4962 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)))))))).
4966 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))))))))).
4967 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))))))))).
4971 Definition l_e_st_eq_landau_n_1to : (forall (n:l_e_st_eq_landau_n_nat), Type).
4972 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)).
4976 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))).
4977 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))).
4981 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)).
4982 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)).
4986 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)).
4987 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)).
4991 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))))))).
4992 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)))))).
4996 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)))))).
4997 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)))))).
5001 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))))).
5002 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)))).
5006 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)))).
5007 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)))).
5011 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)))).
5012 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)).
5016 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))))).
5017 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))).
5021 Definition l_e_st_eq_landau_n_1o : l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_1.
5022 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))).
5026 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).
5027 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).
5031 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).
5032 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).
5036 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).
5037 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))).
5041 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).
5042 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))).
5046 Definition l_e_st_eq_landau_n_2 : l_e_st_eq_landau_n_nat.
5047 exact (l_e_st_eq_landau_n_pl l_e_st_eq_landau_n_1 l_e_st_eq_landau_n_1).
5051 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))).
5052 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)))).
5056 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))).
5057 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))))).
5061 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))).
5062 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))).
5066 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.
5067 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)).
5071 Definition l_e_st_eq_landau_n_1t : l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2.
5072 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)).
5076 Definition l_e_st_eq_landau_n_2t : l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2.
5077 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))).
5081 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).
5082 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).
5086 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).
5087 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).
5091 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)).
5092 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)).
5096 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)).
5097 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))).
5101 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)).
5102 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)).
5106 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)).
5107 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))).
5111 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)).
5112 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)).
5116 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)).
5117 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).
5121 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).
5122 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)))).
5126 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).
5127 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)).
5131 Definition l_e_st_eq_landau_n_pair1type : (forall (alpha:Type), Type).
5132 exact (fun (alpha:Type) => (forall (x:l_e_st_eq_landau_n_1to l_e_st_eq_landau_n_2), alpha)).
5136 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))).
5137 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)))).
5141 Definition l_e_st_eq_landau_n_first1 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), alpha)).
5142 exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => p l_e_st_eq_landau_n_1t)).
5146 Definition l_e_st_eq_landau_n_second1 : (forall (alpha:Type), (forall (p:l_e_st_eq_landau_n_pair1type alpha), alpha)).
5147 exact (fun (alpha:Type) => (fun (p:l_e_st_eq_landau_n_pair1type alpha) => p l_e_st_eq_landau_n_2t)).
5151 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))).
5152 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)))).
5156 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))))).
5157 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)))).
5161 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))).
5162 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))).
5166 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))))).
5167 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)))).
5171 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)))))))).
5172 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))))))).
5176 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)))))))).
5177 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)))))))).
5181 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)))))))).
5182 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)))))))).
5186 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)))))))).
5187 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))))))).
5191 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)))))))).
5192 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)))))))).
5196 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)))))))).
5197 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)))))))).
5201 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))))))).
5202 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))))))).
5206 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))))).
5207 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)))))).
5211 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)).
5212 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))).
5216 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))).
5217 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))).
5221 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))).
5222 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))).
5226 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)).
5227 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))).
5231 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)))).
5232 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))).
5236 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).
5237 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)).
5241 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).
5242 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)).
5246 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).
5247 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)).
5251 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)))).
5252 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)))).
5256 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)))).
5257 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)))).
5261 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)))).
5262 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)))).
5266 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)))).
5267 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))))).
5271 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))))))).
5272 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)))))).
5276 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)))))).
5277 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))))))).
5281 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)))).
5282 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)))))).
5286 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))).
5287 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))).
5291 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))).
5292 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))).
5296 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)))).
5297 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)))).
5301 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)))).
5302 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)))).
5306 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))))))).
5307 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))))).
5311 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)))))).
5312 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)))))).
5316 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))))).
5317 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)))))).
5321 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)))))).
5322 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))))))).
5326 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))))).
5327 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)))))).
5331 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))))).
5332 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))))).
5336 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))))).
5337 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)))))).
5341 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)))))).
5342 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))))).
5346 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)))))).
5347 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)))))).
5351 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)))))))).
5352 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)))))))).
5356 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)))))))).
5357 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)))))))).
5361 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))))))).
5362 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))))))).
5366 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))))))).
5367 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)))))).
5371 Definition l_e_st_eq_landau_n_frac : Type.
5372 exact (l_e_st_eq_landau_n_pair1type l_e_st_eq_landau_n_nat).
5376 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)).
5377 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)).
5381 Definition l_e_st_eq_landau_n_num : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat).
5382 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).
5386 Definition l_e_st_eq_landau_n_den : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat).
5387 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).
5391 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)).
5392 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)).
5396 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)))).
5397 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)).
5401 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)).
5402 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)).
5406 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)))).
5407 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)).
5411 Definition l_e_st_eq_landau_n_1x : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat).
5412 exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_num x).
5416 Definition l_e_st_eq_landau_n_2x : (forall (x:l_e_st_eq_landau_n_frac), l_e_st_eq_landau_n_nat).
5417 exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_den x).
5421 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).
5422 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).
5426 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))).
5427 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).
5431 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))))))).
5432 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))))).
5436 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))))).
5437 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))))).
5441 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))))).
5442 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)))).
5446 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))))).
5447 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)))).
5451 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)))))).
5452 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)))).
5456 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)))).
5457 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)))).
5461 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))))).
5462 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)))).
5466 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))))).
5467 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)))).
5471 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))))))).
5472 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))))))).
5476 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)).
5477 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)).
5481 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)).
5482 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)).
5486 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)).
5487 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)))).
5491 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)))))).
5492 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)))))).
5496 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))))).
5497 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))))).
5501 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)))).
5502 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))))).
5506 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).
5507 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))).
5511 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).
5512 exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_satz37 x).
5516 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))).
5517 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))).
5521 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))).
5522 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))).
5526 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))))))).
5527 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))))))).
5531 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))))).
5532 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))))).
5536 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))))).
5537 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))))).
5541 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))).
5542 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))).
5546 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))).
5547 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))).
5551 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)))))).
5552 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))))).
5556 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)))))).
5557 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)))))).
5561 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)))))).
5562 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))))).
5566 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)))))).
5567 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)))))).
5571 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))).
5572 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))).
5576 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))).
5577 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))).
5581 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)))))))).
5582 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))))).
5586 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)))))))).
5587 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)))))).
5591 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)))))))).
5592 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)))))))).
5596 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)))))))).
5597 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)))))).
5601 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))))).
5602 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)))))).
5606 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)))))))).
5607 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)))))))).
5611 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))))).
5612 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))))).
5616 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))))).
5617 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))))).
5621 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))))).
5622 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)))))).
5626 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))))))).
5627 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)))))))).
5631 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))))))))).
5632 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)))))))))).
5636 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)))).
5637 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))))).
5641 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)).
5642 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))).
5646 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))))).
5647 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))))).
5651 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)))).
5652 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)))).
5656 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)).
5657 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)))).
5661 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)).
5662 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)))).
5666 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)))))).
5667 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))))).
5671 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)))))).
5672 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))))).
5676 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))))).
5677 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)))).
5681 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)))).
5682 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)))).
5686 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))))).
5687 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)))).
5691 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)))).
5692 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)))).
5696 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))).
5697 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)))).
5701 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))).
5702 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)))).
5706 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))).
5707 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)))).
5711 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))).
5712 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))).
5716 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))).
5717 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))).
5721 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)))).
5722 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)))).
5726 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)))).
5727 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)))).
5731 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)))))))))).
5732 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)))))))).
5736 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)))))))))).
5737 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))))))))).
5741 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)))))))))).
5742 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)))))))).
5746 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))))))).
5747 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))))))))).
5751 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))))))).
5752 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))))))).
5756 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))))).
5757 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)))))).
5761 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))))).
5762 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))))).
5766 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))))))).
5767 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)))))))).
5771 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))))))).
5772 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))))))).
5776 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))))).
5777 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)))))).
5781 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))))).
5782 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))))).
5786 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)).
5787 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))).
5791 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)).
5792 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))).
5796 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))).
5797 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))).
5801 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))).
5802 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))).
5806 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))).
5807 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))).
5811 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))).
5812 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))).
5816 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)))).
5817 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)))).
5821 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)))).
5822 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))).
5826 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))).
5827 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))).
5831 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))).
5832 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)))).
5836 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)))).
5837 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)))).
5841 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)))).
5842 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)))).
5846 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))).
5847 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)))).
5851 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))).
5852 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)))).
5856 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)))))))).
5857 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))))))))).
5861 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)))))))).
5862 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))))))))).
5866 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))))))).
5867 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)))))))).
5871 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))))))).
5872 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))))))).
5876 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))))).
5877 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)))))).
5881 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))))).
5882 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))))).
5886 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)))))))).
5887 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))))))))).
5891 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)))))))).
5892 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))))))))).
5896 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))))))).
5897 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)))))))).
5901 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))))))).
5902 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))))))).
5906 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))))).
5907 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)))))).
5911 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))))).
5912 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))))).
5916 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))).
5917 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)))).
5921 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))).
5922 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)))).
5926 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)))))))).
5927 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))))).
5931 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)))))))).
5932 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)))))).
5936 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))))).
5937 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)))))).
5941 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))))).
5942 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))))).
5946 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))))).
5947 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))))))).
5951 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))))).
5952 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)))))).
5956 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))))).
5957 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)))))).
5961 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))))).
5962 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))))))).
5966 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))))).
5967 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))))))).
5971 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))))))).
5972 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)))))))).
5976 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))))))).
5977 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)))))))).
5981 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)))))).
5982 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))))))).
5986 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)))))).
5987 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))))))).
5991 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))))).
5992 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)))))).
5996 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))))).
5997 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))))).
6001 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)))))).
6002 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))))))).
6006 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)))))).
6007 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)))))).
6011 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))))).
6012 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)))))).
6016 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))))).
6017 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))))))).
6021 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))).
6022 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)))).
6026 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).
6027 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)).
6031 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)).
6032 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)).
6036 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)))).
6037 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)))).
6041 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).
6042 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)).
6046 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)).
6047 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)).
6051 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)))))).
6052 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))).
6056 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)))))).
6057 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))).
6061 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))))).
6062 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)))).
6066 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)))))).
6067 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)))).
6071 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)))))).
6072 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)))).
6076 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))).
6077 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)))).
6081 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)))).
6082 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)))).
6086 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))))).
6087 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)))).
6091 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)).
6092 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)))).
6096 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)))))).
6097 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))))).
6101 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))))).
6102 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))))).
6106 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)))))).
6107 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))))).
6111 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)))))).
6112 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)))).
6116 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)))).
6117 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)))).
6121 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))))).
6122 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)))).
6126 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))))).
6127 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)))).
6131 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))))).
6132 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)))).
6136 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)))))).
6137 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)))).
6141 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)))))).
6142 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))))).
6146 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)))))).
6147 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))))).
6151 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))))).
6152 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)))).
6156 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))))).
6157 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)))).
6161 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)))))).
6162 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)))).
6166 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)))).
6167 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)))).
6171 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))))))))).
6172 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)))))).
6176 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))))))))).
6177 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)))))).
6181 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))))))))).
6182 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))))))))).
6186 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))))))))).
6187 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))))))))).
6191 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))))))))).
6192 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)))))))).
6196 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)))))))))).
6197 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))))))).
6201 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))))))))).
6202 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))))))))).
6206 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))))))).
6207 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))))))).
6211 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))))))).
6212 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)))))).
6216 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))))).
6217 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))))).
6221 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))))).
6222 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)))).
6226 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)))).
6227 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)))).
6231 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))))).
6232 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)))).
6236 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))).
6237 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)))).
6241 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))).
6242 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)).
6246 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))))).
6247 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)))))).
6251 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)))))).
6252 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)))).
6256 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)))))).
6257 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)))).
6261 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))))).
6262 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))))).
6266 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)))))).
6267 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)))).
6271 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)))))).
6272 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))))).
6276 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)))))).
6277 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)))).
6281 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))))).
6282 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)))))).
6286 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))))).
6287 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))).
6291 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)))).
6292 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)))).
6296 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)))).
6297 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)))).
6301 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)))))).
6302 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)))))).
6306 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))))).
6307 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))))).
6311 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)))))).
6312 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)))))).
6316 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)))).
6317 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)))).
6321 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)))).
6322 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))).
6326 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))))).
6327 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))))).
6331 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))))).
6332 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))).
6336 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)).
6337 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))).
6341 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))).
6342 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))).
6346 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)))))).
6347 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)))).
6351 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)))))).
6352 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))))).
6356 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)))))).
6357 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))))).
6361 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))))))).
6362 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))))).
6366 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)))))).
6367 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))))).
6371 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)))))).
6372 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))))).
6376 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))))))).
6377 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))))).
6381 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))))).
6382 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))))).
6386 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))))).
6387 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)))).
6391 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))))).
6392 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)))).
6396 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))))).
6397 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)))))).
6401 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))))).
6402 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))))).
6406 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))))).
6407 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)))).
6411 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))))).
6412 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))))).
6416 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))))))).
6417 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))))))).
6421 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))))))).
6422 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))))))).
6426 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))))))).
6427 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))))))).
6431 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))))))).
6432 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))))))).
6436 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)))).
6437 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))).
6441 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))))).
6442 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)))).
6446 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)))).
6447 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)))).
6451 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)))).
6452 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)))).
6456 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)))).
6457 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)))).
6461 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)))).
6462 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))))).
6466 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)))).
6467 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)))))).
6471 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)))).
6472 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))))).
6476 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))))))).
6477 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)))))).
6481 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))))))).
6482 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))))))).
6486 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))))))).
6487 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))))))).
6491 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))))))).
6492 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)))))))).
6496 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))))))).
6497 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))))))).
6501 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))))))).
6502 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)))))))).
6506 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))))))).
6507 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)))))))).
6511 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))))))).
6512 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)))))))).
6516 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))))))))).
6517 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))))))))).
6521 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))))))))).
6522 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))))))))).
6526 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)))))))).
6527 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)))))))).
6531 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)))))))).
6532 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)))))))).
6536 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))))))).
6537 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))))))).
6541 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))))))).
6542 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)))))))).
6546 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))))).
6547 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)))))).
6551 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))))).
6552 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))))).
6556 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))))).
6557 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)))).
6561 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)))))).
6562 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))))))).
6566 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)))).
6567 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))).
6571 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))).
6572 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)))).
6576 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))))).
6577 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)))).
6581 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))).
6582 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))))).
6586 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)))))).
6587 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)))))).
6591 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))).
6592 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))))).
6596 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)))).
6597 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)))).
6601 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))).
6602 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))).
6606 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))).
6607 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))).
6611 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))))).
6612 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)))).
6616 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)))))).
6617 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)))))).
6621 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)).
6622 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)))).
6626 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))))).
6627 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))))).
6631 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))))).
6632 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))))).
6636 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)))))).
6637 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))))).
6641 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)))).
6642 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)))).
6646 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)))).
6647 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)))).
6651 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))))).
6652 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)))).
6656 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))))).
6657 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)))).
6661 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))))).
6662 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)))).
6666 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)))))).
6667 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)))).
6671 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)))))).
6672 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))))).
6676 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)))))).
6677 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))))).
6681 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))))).
6682 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)))).
6686 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))))).
6687 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)))).
6691 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)))))).
6692 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)))).
6696 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)))).
6697 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)))).
6701 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))))))))).
6702 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)))))).
6706 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)))))).
6707 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)))))).
6711 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))))))))).
6712 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)))))))).
6716 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))))))).
6717 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))))))).
6721 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))))))).
6722 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)))))).
6726 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))))).
6727 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))))).
6731 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))))).
6732 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)))).
6736 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))).
6737 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)))).
6741 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))).
6742 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)).
6746 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))))).
6747 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)))))).
6751 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))))).
6752 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))).
6756 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)))).
6757 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)))).
6761 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)))))))).
6762 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))))))).
6766 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)))).
6767 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))))).
6771 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)))).
6772 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)))).
6776 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))))).
6777 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))))).
6781 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))))).
6782 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))))).
6786 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))))).
6787 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))).
6791 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)))).
6792 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)))).
6796 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))))).
6797 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)))).
6801 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))))))).
6802 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)))).
6806 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))))))).
6807 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))))).
6811 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))))).
6812 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))))).
6816 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))))).
6817 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))))).
6821 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))))).
6822 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)))))).
6826 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))))).
6827 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))))).
6831 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))))).
6832 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)))).
6836 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))))).
6837 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))))).
6841 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))))))).
6842 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))))))).
6846 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))))))).
6847 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))))))).
6851 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))))))).
6852 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))))))).
6856 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))))))).
6857 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))))))).
6861 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)))).
6862 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))).
6866 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))))).
6867 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)))).
6871 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)))).
6872 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)))).
6876 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)))).
6877 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)))).
6881 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)))).
6882 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)))).
6886 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)))).
6887 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))))).
6891 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)))).
6892 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)))))).
6896 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)))).
6897 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))))).
6901 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))))))).
6902 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)))))).
6906 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))))))).
6907 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))))))).
6911 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))))))).
6912 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))))))).
6916 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))))))).
6917 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)))))))).
6921 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))))))).
6922 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))))))).
6926 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))))))).
6927 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)))))))).
6931 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))))))).
6932 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)))))))).
6936 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))))))).
6937 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)))))))).
6941 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))))))))).
6942 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))))))))).
6946 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))))))))).
6947 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))))))))).
6951 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)))))))).
6952 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)))))))).
6956 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)))))))).
6957 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)))))))).
6961 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))))))).
6962 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))))))).
6966 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))))))).
6967 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)))))))).
6971 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)))))).
6972 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))))))).
6976 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)).
6977 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)))).
6981 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)).
6982 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))))).
6986 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))).
6987 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))).
6991 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)).
6992 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)).
6996 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).
6997 exact (fun (x:l_e_st_eq_landau_n_frac) => l_e_st_eq_landau_n_refeq x).
7001 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))).
7002 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))).
7006 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))))).
7007 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))))).
7011 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)).
7012 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)).
7016 Definition l_e_st_eq_landau_n_rt_rat : Type.
7017 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).
7021 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)).
7022 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)).
7026 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)).
7027 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))).
7031 Definition l_e_st_eq_landau_n_rt_some : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rat), Prop)), Prop).
7032 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).
7036 Definition l_e_st_eq_landau_n_rt_all : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rat), Prop)), Prop).
7037 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).
7041 Definition l_e_st_eq_landau_n_rt_one : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_rat), Prop)), Prop).
7042 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).
7046 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)).
7047 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)).
7051 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).
7052 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).
7056 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).
7057 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).
7061 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))).
7062 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).
7066 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)))))).
7067 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))))).
7071 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))).
7072 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))).
7076 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)))))).
7077 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)))))))).
7081 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)))).
7082 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)))))).
7086 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))))))).
7087 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))))))))))).
7091 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))))).
7092 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))))))).
7096 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)))))))).
7097 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)))))))))))))).
7101 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)))))).
7102 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)))))))).
7106 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))))))).
7107 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))))))).
7111 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))))))).
7112 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))))))).
7116 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))))))).
7117 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)))))))).
7121 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)))))))).
7122 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)))))))).
7126 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)).
7127 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)).
7131 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))))).
7132 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))))).
7136 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)))))))))).
7137 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))))))))).
7141 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).
7142 exact (fun (x0:l_e_st_eq_landau_n_rt_rat) => l_e_refis l_e_st_eq_landau_n_rt_rat x0).
7146 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))).
7147 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))).
7151 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))))).
7152 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))))).
7156 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)).
7157 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))))).
7161 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)))).
7162 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))))).
7166 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)))))))))))).
7167 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))))))))))).
7171 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)))))))))))).
7172 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))))))))))).
7176 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))))))))))).
7177 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))))))))))).
7181 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))))))))))).
7182 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))))))))))))).
7186 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))))))))).
7187 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))))))))))).
7191 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))))))).
7192 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))))))))).
7196 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))))))).
7197 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))))))).
7201 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)))))))).
7202 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)))))))).
7206 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))))))).
7207 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)))))))).
7211 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))))))).
7212 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))))))).
7216 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))))).
7217 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))))).
7221 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))))).
7222 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))))).
7226 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))))))).
7227 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)))))))).
7231 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)).
7232 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))))).
7236 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)))).
7237 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))))).
7241 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)))))))))))).
7242 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))))))))))).
7246 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)))))))))))).
7247 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))))))))))).
7251 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))))))))))).
7252 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))))))))))).
7256 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))))))))))).
7257 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))))))))))))).
7261 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))))))))).
7262 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))))))))))).
7266 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))))))).
7267 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))))))))).
7271 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))))))).
7272 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))))))).
7276 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)))))))).
7277 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)))))))).
7281 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))))))).
7282 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)))))))).
7286 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))))))).
7287 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))))))).
7291 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))))).
7292 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))))).
7296 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))))).
7297 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))))).
7301 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))))))).
7302 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)))))))).
7306 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))))))).
7307 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)))))).
7311 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)))))))).
7312 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)))))))).
7316 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)))))))).
7317 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)))))))).
7321 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)))))))).
7322 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)))))))).
7326 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))))))).
7327 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))))))).
7331 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))))))).
7332 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)))))).
7336 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)))))))).
7337 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)))))))).
7341 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)))))))).
7342 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)))))))).
7346 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))))))).
7347 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)))))))).
7351 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))))))).
7352 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)))))))).
7356 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))))))).
7357 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))))))).
7361 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))).
7362 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)))))).
7366 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))).
7367 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))).
7371 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))).
7372 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))).
7376 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))))))).
7377 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))))))))).
7381 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))).
7382 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))))))).
7386 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))))))).
7387 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))))))))).
7391 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))).
7392 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))))))).
7396 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)).
7397 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))).
7401 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))).
7402 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))).
7406 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))).
7407 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))).
7411 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))))))).
7412 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))))))))).
7416 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))))))).
7417 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))))))))).
7421 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))))).
7422 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))))).
7426 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))))).
7427 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))))).
7431 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))))))).
7432 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)))))))).
7436 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)).
7437 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))).
7441 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))).
7442 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))).
7446 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))).
7447 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))).
7451 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))))))).
7452 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))))))))).
7456 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))))))).
7457 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))))))))).
7461 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))))).
7462 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))))).
7466 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))))).
7467 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))))).
7471 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))))))).
7472 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)))))))).
7476 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)))).
7477 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)))).
7481 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)))).
7482 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))).
7486 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))).
7487 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))).
7491 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))).
7492 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)))).
7496 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)))).
7497 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)))).
7501 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)))).
7502 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)))).
7506 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))).
7507 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)))).
7511 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))).
7512 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)))).
7516 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))))))).
7517 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))))))))).
7521 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))).
7522 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))))))).
7526 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))))))).
7527 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))))))))).
7531 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))).
7532 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))))))).
7536 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))))))))))).
7537 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))))))))))))).
7541 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))))).
7542 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))))))))))).
7546 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))))).
7547 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))))).
7551 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))))).
7552 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))))))).
7556 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))))))))))).
7557 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))))))))))))).
7561 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))))).
7562 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))))))))))).
7566 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))))))))))).
7567 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))))))))))))).
7571 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))))).
7572 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))))))))))).
7576 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))))).
7577 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))))))).
7581 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))))).
7582 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))))))).
7586 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))))))))))).
7587 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))))))))))))).
7591 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))))).
7592 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))))))))))).
7596 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))))).
7597 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))))).
7601 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))))).
7602 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))))))).
7606 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)))))).
7607 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)))))).
7611 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)))).
7612 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))))).
7616 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)).
7617 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))).
7621 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)))))).
7622 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)))))).
7626 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)))).
7627 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))))).
7631 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)).
7632 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))).
7636 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)))))))))).
7637 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)))))))))).
7641 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))))))))).
7642 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)))))))))).
7646 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)))))))))).
7647 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)))))))))).
7651 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))))))))))).
7652 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)))))))))).
7656 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))))))))).
7657 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))))))))).
7661 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))))).
7662 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))))))).
7666 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)).
7667 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))).
7671 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))))))).
7672 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))))))).
7676 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.
7677 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)))))).
7681 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)).
7682 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)).
7686 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))))))).
7687 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)))))).
7691 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)))))))).
7692 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))))))).
7696 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))))).
7697 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)))).
7701 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))))).
7702 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)))).
7706 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))))))).
7707 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))))))).
7711 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))))))).
7712 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))))))).
7716 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))).
7717 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)))))).
7721 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))).
7722 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)).
7726 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))))))))))).
7727 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))))))))).
7731 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)))))))))))).
7732 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)))))))))).
7736 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))))))))))).
7737 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)))))))))).
7741 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))))).
7742 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))))))))).
7746 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))))).
7747 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))).
7751 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)))).
7752 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)))).
7756 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)))))).
7757 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))))))).
7761 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)).
7762 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)))))).
7766 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))).
7767 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))).
7771 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))))))))))).
7772 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)))))))))))).
7776 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))))).
7777 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)))))))))).
7781 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))))))))))).
7782 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)))))))))))).
7786 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))))).
7787 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)))))))))).
7791 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))))))))))).
7792 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)))))))))))).
7796 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))))).
7797 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)))))))))).
7801 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))))))))))).
7802 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)))))))))))).
7806 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))))).
7807 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)))))))))).
7811 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))))).
7812 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)))).
7816 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))))).
7817 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)))).
7821 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))))).
7822 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)))))).
7826 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))))).
7827 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))))).
7831 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))))).
7832 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)))).
7836 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))))).
7837 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))))).
7841 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)))))))))).
7842 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)))))))))))).
7846 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)))).
7847 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)))))))))).
7851 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)))))))))).
7852 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)))))))))))).
7856 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)))).
7857 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)))))))))).
7861 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)))))))))).
7862 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)))))))))))).
7866 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)))).
7867 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)))))))))).
7871 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)))).
7872 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)))))).
7876 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))))))))))))))).
7877 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)))))))))))))))).
7881 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))))))).
7882 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)))))))))))))).
7886 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))))))).
7887 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)))))))).
7891 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))))))))))))))).
7892 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)))))))))))))))).
7896 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))))))).
7897 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)))))))))))))).
7901 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))))))))))))))).
7902 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)))))))))))))))).
7906 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))))))).
7907 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)))))))))))))).
7911 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))))))).
7912 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)))))))).
7916 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))))))).
7917 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)))))))).
7921 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))))))))))))))).
7922 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)))))))))))))))).
7926 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))))))).
7927 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)))))))))))))).
7931 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))))))).
7932 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)))))))).
7936 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))))).
7937 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)))))).
7941 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)))).
7942 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))))).
7946 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))))).
7947 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)))).
7951 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))))))))).
7952 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))))))))).
7956 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)))))))))).
7957 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)))))))))).
7961 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)))))))).
7962 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))))))))).
7966 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)))).
7967 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))))))).
7971 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)))))))))))))).
7972 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)))))))))))))))).
7976 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)))))).
7977 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)))))))))))))).
7981 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))).
7982 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)))))).
7986 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)))).
7987 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)))).
7991 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))).
7992 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)))).
7996 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))).
7997 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)))).
8001 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))))).
8002 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)))).
8006 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))).
8007 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)))).
8011 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)))).
8012 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)))).
8016 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)))))).
8017 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)))))).
8021 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)).
8022 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))).
8026 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))))))).
8027 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))))))).
8031 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.
8032 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)))))).
8036 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)).
8037 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)).
8041 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))))))).
8042 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)))))).
8046 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)))))))).
8047 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))))))).
8051 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))))).
8052 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)))).
8056 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))))).
8057 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)))).
8061 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))))))).
8062 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))))))).
8066 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))))))).
8067 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))))))).
8071 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))).
8072 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)))))).
8076 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))).
8077 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)).
8081 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))))))))))).
8082 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))))))))).
8086 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)))))))))))).
8087 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)))))))))).
8091 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))))))))))).
8092 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)))))))))).
8096 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))))).
8097 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))))))))).
8101 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))))).
8102 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))).
8106 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)))).
8107 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)))).
8111 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)))))))))))).
8112 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)))))))))).
8116 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)))))))))))).
8117 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)))))))))).
8121 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))))))))))).
8122 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)))))))))).
8126 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))))).
8127 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))))))))).
8131 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))))).
8132 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))))).
8136 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))))).
8137 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))).
8141 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)))).
8142 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)))).
8146 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))))).
8147 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)))).
8151 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))))))))))).
8152 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)))))))))))).
8156 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))))).
8157 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)))))))))).
8161 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))))))))))).
8162 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)))))))))))).
8166 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))))).
8167 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)))))))))).
8171 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))))))))))).
8172 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)))))))))))).
8176 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))))).
8177 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)))))))))).
8181 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))))).
8182 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)))).
8186 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))))).
8187 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)))))).
8191 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))))).
8192 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))))).
8196 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))))).
8197 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)))).
8201 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))))).
8202 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))))).
8206 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)))))))))).
8207 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)))))))))))).
8211 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)))).
8212 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)))))))))).
8216 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)))))))))).
8217 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)))))))))))).
8221 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)))).
8222 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)))))))))).
8226 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)))))))))).
8227 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)))))))))))).
8231 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)))).
8232 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)))))))))).
8236 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)))).
8237 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)))))).
8241 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))))))))))))))).
8242 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)))))))))))))))).
8246 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))))))).
8247 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)))))))))))))).
8251 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))))))).
8252 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)))))))).
8256 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))))))))))))))).
8257 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)))))))))))))))).
8261 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))))))).
8262 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)))))))))))))).
8266 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))))))))))))))).
8267 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)))))))))))))))).
8271 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))))))).
8272 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)))))))))))))).
8276 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))))))).
8277 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)))))))).
8281 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))))))).
8282 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)))))))).
8286 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))))))))))))))).
8287 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)))))))))))))))).
8291 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))))))).
8292 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)))))))))))))).
8296 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))))))).
8297 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)))))))).
8301 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)))))))).
8302 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)))))))).
8306 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))))))))).
8307 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))))))))).
8311 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))))))).
8312 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)))))))).
8316 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))).
8317 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)))))).
8321 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)))))))))))))).
8322 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)))))))))))))))).
8326 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)))))).
8327 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)))))))))))))).
8331 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))).
8332 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)))))).
8336 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))).
8337 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))).
8341 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)).
8342 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))).
8346 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))))).
8347 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))).
8351 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))).
8352 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))).
8356 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))).
8357 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)))).
8361 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))).
8362 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))).
8366 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)))).
8367 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))).
8371 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)))).
8372 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)))).
8376 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)))).
8377 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))).
8381 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)).
8382 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))).
8386 Definition l_e_st_eq_landau_n_rt_natrt : (forall (x0:l_e_st_eq_landau_n_rt_rat), Prop).
8387 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)).
8391 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))))))).
8392 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)))))))).
8396 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)).
8397 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)))))).
8401 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))).
8402 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)).
8406 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)).
8407 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))).
8411 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))).
8412 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))).
8416 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)))))).
8417 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))))).
8421 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))))).
8422 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)))))).
8426 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).
8427 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)).
8431 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)).
8432 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))).
8436 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)))).
8437 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))).
8441 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))).
8442 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))).
8446 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)))).
8447 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)))).
8451 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))).
8452 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))).
8456 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))).
8457 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)).
8461 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))).
8462 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)))).
8466 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)))))).
8467 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)))))).
8471 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))))).
8472 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))))).
8476 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)))))).
8477 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)))))).
8481 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))))).
8482 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))))).
8486 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)))))).
8487 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)))))).
8491 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)))))))))).
8492 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))))))))).
8496 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))))))))).
8497 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))))))))).
8501 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)))))))).
8502 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))))))))).
8506 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)))))))).
8507 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)))))))).
8511 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)))))).
8512 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))))))).
8516 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)))).
8517 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))).
8521 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)))).
8522 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))).
8526 Definition l_e_st_eq_landau_n_rt_nt_natt : Type.
8527 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)).
8531 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)).
8532 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)).
8536 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)).
8537 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)).
8541 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)).
8542 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))).
8546 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).
8547 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).
8551 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).
8552 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).
8556 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).
8557 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).
8561 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)).
8562 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)).
8566 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).
8567 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).
8571 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)).
8572 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).
8576 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)))))).
8577 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))))).
8581 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))))).
8582 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))))).
8586 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)))).
8587 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))).
8591 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))).
8592 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))).
8596 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)))).
8597 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)).
8601 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))).
8602 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).
8606 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).
8607 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)).
8611 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)))).
8612 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)))).
8616 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))).
8617 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)))).
8621 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).
8622 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)).
8626 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)))).
8627 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)))).
8631 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))).
8632 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)))).
8636 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))).
8637 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)).
8641 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))).
8642 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)).
8646 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))).
8647 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)).
8651 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))).
8652 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)).
8656 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))).
8657 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)).
8661 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))).
8662 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)).
8666 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).
8667 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)).
8671 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).
8672 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)).
8676 Definition l_e_st_eq_landau_n_rt_nt_1t : l_e_st_eq_landau_n_rt_nt_natt.
8677 exact (l_e_st_eq_landau_n_rt_nt_ntofn l_e_st_eq_landau_n_1).
8681 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).
8682 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))).
8686 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)).
8687 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)).
8691 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).
8692 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)).
8696 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))))).
8697 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))).
8701 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))).
8702 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))))).
8706 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).
8707 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).
8711 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).
8712 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))).
8716 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)))).
8717 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)))).
8721 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)))))).
8722 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)))))).
8726 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))))).
8727 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)))))).
8731 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)))))).
8732 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)))))).
8736 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)))).
8737 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))))).
8741 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)))).
8742 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))))).
8746 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).
8747 exact (fun (x:l_e_st_eq_landau_n_rt_nt_natt) => l_e_st_eq_landau_n_rt_nt_satz113a x).
8751 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))).
8752 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))).
8756 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)))).
8757 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)))).
8761 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)))).
8762 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)))).
8766 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))))).
8767 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)))).
8771 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)))).
8772 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)))).
8776 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)))).
8777 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)))).
8781 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))))).
8782 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)))))))).
8786 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))))).
8787 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)))))).
8791 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))))).
8792 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)))))))).
8796 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))))).
8797 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)))))).
8801 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)))))).
8802 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)))))))).
8806 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))))))).
8807 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)))))).
8811 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)))))))).
8812 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))))))))).
8816 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)))))).
8817 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)))))).
8821 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))))))).
8822 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))))))).
8826 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)))))))))).
8827 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)))))))).
8831 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))))))))))).
8832 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)))))))).
8836 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)))))))))).
8837 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)))))))).
8841 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))))))))))).
8842 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)))))))).
8846 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))))))))).
8847 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))))))))).
8851 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))))))))).
8852 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)))))))).
8856 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))))))).
8857 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))))))).
8861 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))))))).
8862 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))))))).
8866 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))))))))).
8867 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)))))))))).
8871 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))))))))).
8872 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))))))))).
8876 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))))))))).
8877 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))))))))).
8881 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))))))))))))).
8882 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)))))))))).
8886 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)))))))))))).
8887 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)))))))))).
8891 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)))))))))).
8892 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))))))))))).
8896 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)))))))).
8897 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))))))))).
8901 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))))).
8902 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)))))))).
8906 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))))).
8907 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)))))).
8911 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)))))).
8912 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)))))))).
8916 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)))))))).
8917 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)))))).
8921 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))))))).
8922 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))))))).
8926 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))))))).
8927 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)))))))).
8931 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)))))).
8932 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)))))).
8936 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))))))).
8937 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))))))).
8941 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)))))))))).
8942 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)))))))))).
8946 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)))))))))).
8947 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)))))))).
8951 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)))))))))).
8952 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)))))))))).
8956 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))))))))).
8957 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))))))))).
8961 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))))))).
8962 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))))))).
8966 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))))))).
8967 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))))))).
8971 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))))))).
8972 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))))))).
8976 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))))).
8977 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)))))).
8981 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))))))))).
8982 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))))).
8986 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)).
8987 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)))).
8991 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))))).
8992 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))).
8996 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))))).
8997 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)))).
9001 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))))).
9002 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))).
9006 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))).
9007 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))).
9011 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)))).
9012 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)))).
9016 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)))).
9017 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))).
9021 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))))).
9022 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)))).
9026 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)))))).
9027 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)))).
9031 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)))))).
9032 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)))).
9036 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))))).
9037 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))))).
9041 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))).
9042 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)))).
9046 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))))).
9047 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))))).
9051 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)))))).
9052 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))))).
9056 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)))).
9057 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))))).
9061 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)).
9062 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)).
9066 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)).
9067 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))).
9071 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)).
9072 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))).
9076 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)).
9077 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))).
9081 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)).
9082 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))).
9086 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)).
9087 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))).
9091 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)))).
9092 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)))).
9096 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))))))).
9097 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)))))).
9101 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))))).
9102 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))))))).
9106 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))).
9107 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))))).
9111 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)))).
9112 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)))).
9116 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)))).
9117 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)))).
9121 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))).
9122 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))).
9126 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))).
9127 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))).
9131 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))))).
9132 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)))))).
9136 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))).
9137 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))))).
9141 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))).
9142 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))).
9146 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))).
9147 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))).
9151 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)))).
9152 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)))).
9156 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)))).
9157 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)))).
9161 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))).
9162 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))).
9166 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)))).
9167 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)))).
9171 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)))).
9172 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)))).
9176 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))).
9177 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))).
9181 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)))).
9182 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)))).
9186 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)))).
9187 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)))).
9191 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))).
9192 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))).
9196 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))).
9197 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))).
9201 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))))).
9202 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))).
9206 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)).
9207 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))).
9211 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)))).
9212 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))).
9216 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)))).
9217 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)))).
9221 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)))).
9222 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))).
9226 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))).
9227 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)))).
9231 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)).
9232 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))).
9236 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)))).
9237 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))).
9241 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)))).
9242 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)))).
9246 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)))).
9247 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))).
9251 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))).
9252 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)))).
9256 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)).
9257 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))).
9261 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)))).
9262 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)))).
9266 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))).
9267 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)))).
9271 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)).
9272 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))).
9276 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)))).
9277 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)))).
9281 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))).
9282 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)))).
9286 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)))))).
9287 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)))))).
9291 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))))).
9292 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)))))).
9296 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)))))))).
9297 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))))))).
9301 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)))))))).
9302 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))))))).
9306 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))))))).
9307 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))))))).
9311 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)).
9312 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)))).
9316 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)).
9317 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))).
9321 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))).
9322 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)))).
9326 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))))).
9327 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))))).
9331 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))))).
9332 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))))).
9336 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))).
9337 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)))).
9341 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))).
9342 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))).
9346 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)))).
9347 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)))).
9351 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)))))).
9352 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))))))).
9356 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))))))).
9357 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))))))).
9361 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))))))).
9362 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))))))).
9366 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)))))).
9367 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))))))).
9371 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)))))).
9372 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)))))).
9376 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))))).
9377 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))))).
9381 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))))).
9382 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)))).
9386 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))))).
9387 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))))).
9391 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))))).
9392 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))))).
9396 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))).
9397 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)))).
9401 Definition l_e_st_eq_landau_n_rt_1rt : l_e_st_eq_landau_n_rt_rat.
9402 exact (l_e_st_eq_landau_n_rt_rtofn l_e_st_eq_landau_n_1).
9406 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))).
9407 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))))).
9411 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))).
9412 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)))).
9416 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).
9417 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))).
9421 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)).
9422 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)).
9426 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).
9427 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)).
9431 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)).
9432 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)).
9436 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)).
9437 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))).
9441 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))).
9442 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)).
9446 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))).
9447 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)))).
9451 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)).
9452 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))).
9456 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)).
9457 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))).
9461 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)))).
9462 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))).
9466 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)).
9467 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))).
9471 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))).
9472 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))).
9476 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))))).
9477 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))))).
9481 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)))).
9482 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)))).
9486 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)))).
9487 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))).
9491 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)))).
9492 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))).
9496 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)))))).
9497 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)))))).
9501 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)))))).
9502 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)))))).
9506 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))))))).
9507 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)))))).
9511 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)))))).
9512 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)))))))).
9516 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))))))))).
9517 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))))))))).
9521 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)))))))).
9522 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)))))))).
9526 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)))))))).
9527 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)))))))).
9531 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)))))))).
9532 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))))))))).
9536 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))))))).
9537 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))))))).
9541 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)))))).
9542 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)))))))).
9546 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)))))).
9547 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))))))).
9551 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))))))).
9552 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))))))).
9556 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))))).
9557 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)))))).
9561 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))).
9562 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)))).
9566 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))))))).
9567 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))))))).
9571 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)))))))).
9572 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))))))).
9576 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)))))).
9577 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)))))).
9581 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)))).
9582 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)))).
9586 Definition l_e_st_eq_landau_n_rt_cutprop1a : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop).
9587 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).
9591 Definition l_e_st_eq_landau_n_rt_cutprop1b : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop).
9592 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))).
9596 Definition l_e_st_eq_landau_n_rt_cutprop1 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop).
9597 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)).
9601 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)).
9602 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)))).
9606 Definition l_e_st_eq_landau_n_rt_cutprop2 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop).
9607 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))).
9611 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))).
9612 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)))).
9616 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)).
9617 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))).
9621 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)).
9622 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))).
9626 Definition l_e_st_eq_landau_n_rt_cutprop3 : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop).
9627 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))).
9631 Definition l_e_st_eq_landau_n_rt_cutprop : (forall (s:l_e_st_set l_e_st_eq_landau_n_rt_rat), Prop).
9632 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)).
9636 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))).
9637 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)))).
9641 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)).
9642 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))).
9646 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)).
9647 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))).
9651 Definition l_e_st_eq_landau_n_rt_cut : Type.
9652 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)).
9656 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).
9657 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).
9661 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)).
9662 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))).
9666 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)).
9667 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)))).
9671 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)).
9672 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).
9676 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)).
9677 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)).
9681 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)).
9682 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)).
9686 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)).
9687 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)).
9691 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)).
9692 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)).
9696 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)).
9697 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)).
9701 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))).
9702 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))).
9706 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)).
9707 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)).
9711 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))).
9712 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))).
9716 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))).
9717 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)))).
9721 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))))).
9722 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)))))).
9726 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))))).
9727 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)))))).
9731 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)))).
9732 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))).
9736 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)))).
9737 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))).
9741 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))))).
9742 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)))).
9746 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))))))).
9747 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))))))).
9751 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)))))))).
9752 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))))))).
9756 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))))))).
9757 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)))))))).
9761 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))))))).
9762 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)))))))).
9766 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))))).
9767 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))))))).
9771 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))))).
9772 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)))))).
9776 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)).
9777 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)).
9781 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))))))).
9782 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)))))))).
9786 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)).
9787 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)).
9791 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)).
9792 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))).
9796 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)))).
9797 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))).
9801 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))))).
9802 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))))).
9806 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))).
9807 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))).
9811 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)))).
9812 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))))).
9816 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)).
9817 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)).
9821 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)))).
9822 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)))).
9826 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)))).
9827 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)))).
9831 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))))))).
9832 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))))))).
9836 Definition l_e_st_eq_landau_n_rt_rp_all : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_cut), Prop)), Prop).
9837 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).
9841 Definition l_e_st_eq_landau_n_rt_rp_some : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_cut), Prop)), Prop).
9842 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).
9846 Definition l_e_st_eq_landau_n_rt_rp_one : (forall (p:(forall (x:l_e_st_eq_landau_n_rt_cut), Prop)), Prop).
9847 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).
9851 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).
9852 exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_e_refis l_e_st_eq_landau_n_rt_cut ksi).
9856 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))).
9857 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))).
9861 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))))).
9862 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))))).
9866 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)))).
9867 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))).
9871 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))))).
9872 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)))))).
9876 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))))).
9877 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)))))).
9881 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)))).
9882 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))).
9886 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))))).
9887 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)))))).
9891 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))))).
9892 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)))))).
9896 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))))))).
9897 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)))))).
9901 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))))))).
9902 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))))))).
9906 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)))))).
9907 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))))))).
9911 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))))))).
9912 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))))))).
9916 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)))))).
9917 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))))))).
9921 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)))))).
9922 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))))))).
9926 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)))))).
9927 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))))))).
9931 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)).
9932 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)))))).
9936 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)))))).
9937 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)))).
9941 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)))))).
9942 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)))))).
9946 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)))))).
9947 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)))))).
9951 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))))))).
9952 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))))))).
9956 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))))))).
9957 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))))))).
9961 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))))))).
9962 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))))))).
9966 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))))))).
9967 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))))))).
9971 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))))))).
9972 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))))))).
9976 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))))).
9977 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)))))).
9981 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))))).
9982 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)))).
9986 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)))).
9987 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)))).
9991 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)).
9992 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))).
9996 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))))))).
9997 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)))))))).
10000 (* constant 2011 *)
10001 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)).
10002 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)))).
10005 (* constant 2012 *)
10006 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))))))).
10007 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))))))).
10010 (* constant 2013 *)
10011 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))))))).
10012 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))))))).
10015 (* constant 2014 *)
10016 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))))))).
10017 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)))))))).
10020 (* constant 2015 *)
10021 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))))).
10022 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))))))).
10025 (* constant 2016 *)
10026 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)).
10027 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)))).
10030 (* constant 2017 *)
10031 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))))))).
10032 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))))))).
10035 (* constant 2018 *)
10036 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))))))).
10037 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))))))).
10040 (* constant 2019 *)
10041 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))))))).
10042 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)))))))).
10045 (* constant 2020 *)
10046 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))))).
10047 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))))))).
10050 (* constant 2021 *)
10051 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))))))).
10052 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)))))).
10055 (* constant 2022 *)
10056 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)))))).
10057 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))))))).
10060 (* constant 2023 *)
10061 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))).
10062 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)))))).
10065 (* constant 2024 *)
10066 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))))))).
10067 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)))))).
10070 (* constant 2025 *)
10071 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)))))).
10072 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))))))).
10075 (* constant 2026 *)
10076 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))).
10077 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)))))).
10080 (* constant 2027 *)
10081 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)))))))).
10082 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)))))).
10085 (* constant 2028 *)
10086 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))))))).
10087 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))))))).
10090 (* constant 2029 *)
10091 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)))).
10092 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)))))).
10095 (* constant 2030 *)
10096 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))).
10097 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))).
10100 (* constant 2031 *)
10101 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)))))))).
10102 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)))))).
10105 (* constant 2032 *)
10106 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))))))).
10107 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))))))).
10110 (* constant 2033 *)
10111 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)))).
10112 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)))))).
10115 (* constant 2034 *)
10116 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))).
10117 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))).
10120 (* constant 2035 *)
10121 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)))))))))).
10122 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)))))))))).
10125 (* constant 2036 *)
10126 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)))))))))).
10127 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)))))))))).
10130 (* constant 2037 *)
10131 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)))))))))).
10132 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)))))))))))).
10135 (* constant 2038 *)
10136 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))))))).
10137 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)))))))))).
10140 (* constant 2039 *)
10141 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)))).
10142 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))))))).
10145 (* constant 2040 *)
10146 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)))).
10147 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)))).
10150 (* constant 2041 *)
10151 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))).
10152 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))).
10155 (* constant 2042 *)
10156 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))).
10157 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))).
10160 (* constant 2043 *)
10161 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)))).
10162 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))).
10165 (* constant 2044 *)
10166 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))))).
10167 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)))).
10170 (* constant 2045 *)
10171 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)))).
10172 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)))).
10175 (* constant 2046 *)
10176 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)))).
10177 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)))).
10180 (* constant 2047 *)
10181 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)))).
10182 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)))).
10185 (* constant 2048 *)
10186 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))).
10187 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))).
10190 (* constant 2049 *)
10191 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))).
10192 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))).
10195 (* constant 2050 *)
10196 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))).
10197 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))).
10200 (* constant 2051 *)
10201 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))).
10202 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))).
10205 (* constant 2052 *)
10206 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)).
10207 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))).
10210 (* constant 2053 *)
10211 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)).
10212 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))).
10215 (* constant 2054 *)
10216 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))).
10217 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)))).
10220 (* constant 2055 *)
10221 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))).
10222 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)))).
10225 (* constant 2056 *)
10226 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))))).
10227 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))))).
10230 (* constant 2057 *)
10231 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))))).
10232 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))))).
10235 (* constant 2058 *)
10236 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))))).
10237 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))))).
10240 (* constant 2059 *)
10241 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))))).
10242 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))))).
10245 (* constant 2060 *)
10246 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))))).
10247 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))))).
10250 (* constant 2061 *)
10251 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))))).
10252 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))))).
10255 (* constant 2062 *)
10256 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))))).
10257 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))))).
10260 (* constant 2063 *)
10261 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))))).
10262 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))))).
10265 (* constant 2064 *)
10266 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))).
10267 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))).
10270 (* constant 2065 *)
10271 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))).
10272 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))).
10275 (* constant 2066 *)
10276 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))).
10277 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))).
10280 (* constant 2067 *)
10281 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))).
10282 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))).
10285 (* constant 2068 *)
10286 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))))))).
10287 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)))))))).
10290 (* constant 2069 *)
10291 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))))))).
10292 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)))))))).
10295 (* constant 2070 *)
10296 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))))))).
10297 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)))))))).
10300 (* constant 2071 *)
10301 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))))))).
10302 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)))))))).
10305 (* constant 2072 *)
10306 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)))).
10307 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)))).
10310 (* constant 2073 *)
10311 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)))).
10312 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))).
10315 (* constant 2074 *)
10316 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))).
10317 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))).
10320 (* constant 2075 *)
10321 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))).
10322 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)))).
10325 (* constant 2076 *)
10326 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)))).
10327 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)))).
10330 (* constant 2077 *)
10331 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)))).
10332 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)))).
10335 (* constant 2078 *)
10336 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))).
10337 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)))).
10340 (* constant 2079 *)
10341 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))).
10342 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)))).
10345 (* constant 2080 *)
10346 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))))))))))).
10347 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))))))))))).
10350 (* constant 2081 *)
10351 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))))))))))).
10352 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)))))))))))).
10355 (* constant 2082 *)
10356 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)))))))))))).
10357 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))))))))))).
10360 (* constant 2083 *)
10361 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))))))))))).
10362 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)))))))))))).
10365 (* constant 2084 *)
10366 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)))))))).
10367 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))))))))))).
10370 (* constant 2085 *)
10371 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))))).
10372 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)))))))).
10375 (* constant 2086 *)
10376 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))))).
10377 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))))).
10380 (* constant 2087 *)
10381 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))))).
10382 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))))))).
10385 (* constant 2088 *)
10386 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))))).
10387 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)))))).
10390 (* constant 2089 *)
10391 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))))).
10392 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)))))).
10395 (* constant 2090 *)
10396 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))))).
10397 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))))))).
10400 (* constant 2091 *)
10401 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))))).
10402 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))))))).
10405 (* constant 2092 *)
10406 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))))))).
10407 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)))))))).
10410 (* constant 2093 *)
10411 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))))))).
10412 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)))))))).
10415 (* constant 2094 *)
10416 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)))))).
10417 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))))))).
10420 (* constant 2095 *)
10421 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)))))).
10422 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))))))).
10425 (* constant 2096 *)
10426 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))))).
10427 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)))))).
10430 (* constant 2097 *)
10431 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))))).
10432 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))))).
10435 (* constant 2098 *)
10436 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))))).
10437 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))))))).
10440 (* constant 2099 *)
10441 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))))).
10442 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))))))).
10445 (* constant 2100 *)
10446 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))).
10447 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))))).
10450 (* constant 2101 *)
10451 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)).
10452 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))).
10455 (* constant 2102 *)
10456 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)))))))).
10457 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)))))))).
10460 (* constant 2103 *)
10461 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))))))))).
10462 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))))))))).
10465 (* constant 2104 *)
10466 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)))))))).
10467 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))))))))).
10470 (* constant 2105 *)
10471 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))))))))).
10472 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))))))))).
10475 (* constant 2106 *)
10476 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)))))).
10477 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)))))).
10480 (* constant 2107 *)
10481 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)))))))))).
10482 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)))))))))).
10485 (* constant 2108 *)
10486 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)))))))))).
10487 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)))))))))).
10490 (* constant 2109 *)
10491 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))))))))))).
10492 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)))))))))).
10495 (* constant 2110 *)
10496 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)))))))))).
10497 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))))))))))).
10500 (* constant 2111 *)
10501 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)))))))).
10502 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)))))))))).
10505 (* constant 2112 *)
10506 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)))))).
10507 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)))))))).
10510 (* constant 2113 *)
10511 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))))))))))))).
10512 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))))))))))))).
10515 (* constant 2114 *)
10516 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))))))))))))).
10517 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))))))))))))).
10520 (* constant 2115 *)
10521 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)))))))))))))).
10522 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))))))))))))))).
10525 (* constant 2116 *)
10526 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)))))))))))))).
10527 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)))))))))))))).
10530 (* constant 2117 *)
10531 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))))))))).
10532 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))))))))))))).
10535 (* constant 2118 *)
10536 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)))))))).
10537 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))))))).
10540 (* constant 2119 *)
10541 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)))))))))))).
10542 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))))))))))).
10545 (* constant 2120 *)
10546 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))))))))))).
10547 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)))))))))))).
10550 (* constant 2121 *)
10551 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))))))))))))).
10552 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)))))))))))).
10555 (* constant 2122 *)
10556 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))))))))))).
10557 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)))))))))))).
10560 (* constant 2123 *)
10561 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))))))))))).
10562 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))))))))))))).
10565 (* constant 2124 *)
10566 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))))))))))).
10567 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))))))))))))).
10570 (* constant 2125 *)
10571 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))))))))))))).
10572 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)))))))))))).
10575 (* constant 2126 *)
10576 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))))))))))))).
10577 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)))))))))))).
10580 (* constant 2127 *)
10581 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))))))))))).
10582 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))))))))))))).
10585 (* constant 2128 *)
10586 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)))))))))))))).
10587 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)))))))))))).
10590 (* constant 2129 *)
10591 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)))))))))))).
10592 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)))))))))))).
10595 (* constant 2130 *)
10596 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))))))).
10597 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))))))))))).
10600 (* constant 2131 *)
10601 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))))))))))))).
10602 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)))))))))))))).
10605 (* constant 2132 *)
10606 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))))))))))))).
10607 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))))))))))))).
10610 (* constant 2133 *)
10611 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)))))))))))).
10612 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))))))))))))).
10615 (* constant 2134 *)
10616 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))))))))))))).
10617 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))))))))))))).
10620 (* constant 2135 *)
10621 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)))))))))))))).
10622 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))))))))))))).
10625 (* constant 2136 *)
10626 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))))))))))).
10627 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)))))))))))).
10630 (* constant 2137 *)
10631 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)))))).
10632 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))))))))).
10635 (* constant 2138 *)
10636 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))))))))))).
10637 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)))))))))))).
10640 (* constant 2139 *)
10641 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))))))))).
10642 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)))))))))).
10645 (* constant 2140 *)
10646 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))))))).
10647 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)))))))).
10650 (* constant 2141 *)
10651 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))))).
10652 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)))))).
10655 (* constant 2142 *)
10656 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))).
10657 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)))).
10660 (* constant 2143 *)
10661 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)).
10662 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))).
10665 (* constant 2144 *)
10666 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)))))))).
10667 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))))))))).
10670 (* constant 2145 *)
10671 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)))))))))).
10672 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)))))))).
10675 (* constant 2146 *)
10676 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)))))))).
10677 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))))))))).
10680 (* constant 2147 *)
10681 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))))))).
10682 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)))))).
10685 (* constant 2148 *)
10686 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)))))).
10687 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)))))).
10690 (* constant 2149 *)
10691 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))))).
10692 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)))).
10695 (* constant 2150 *)
10696 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))))).
10697 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)))).
10700 (* constant 2151 *)
10701 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))))))).
10702 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))))))).
10705 (* constant 2152 *)
10706 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)))))))))).
10707 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)))))))))).
10710 (* constant 2153 *)
10711 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))))))))).
10712 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)))))))))).
10715 (* constant 2154 *)
10716 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)))).
10717 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))))))))).
10720 (* constant 2155 *)
10721 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))).
10722 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)))).
10725 (* constant 2156 *)
10726 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))).
10727 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)).
10730 (* constant 2157 *)
10731 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))))))))))))))))).
10732 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)))))))))))))))).
10735 (* constant 2158 *)
10736 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)))))))))))))))).
10737 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))))))))))))))))).
10740 (* constant 2159 *)
10741 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))))))))))))))).
10742 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)))))))))))))))).
10745 (* constant 2160 *)
10746 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)))))))))).
10747 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))))))))))))))).
10750 (* constant 2161 *)
10751 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))))).
10752 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)))))))))).
10755 (* constant 2162 *)
10756 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)))))))))))))))).
10757 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)))))))))))))))).
10760 (* constant 2163 *)
10761 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)))))))))))))))).
10762 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))))))))))))))))).
10765 (* constant 2164 *)
10766 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))))))))))))))).
10767 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)))))))))))))))).
10770 (* constant 2165 *)
10771 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)))))))))).
10772 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))))))))))))))).
10775 (* constant 2166 *)
10776 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))))).
10777 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)))))))))).
10780 (* constant 2167 *)
10781 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))))).
10782 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))))).
10785 (* constant 2168 *)
10786 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))))).
10787 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))).
10790 (* constant 2169 *)
10791 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)))).
10792 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)))).
10795 (* constant 2170 *)
10796 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)))).
10797 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))))).
10800 (* constant 2171 *)
10801 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))))).
10802 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)))))).
10805 (* constant 2172 *)
10806 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))))).
10807 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))))).
10810 (* constant 2173 *)
10811 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)))).
10812 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))))).
10815 (* constant 2174 *)
10816 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)).
10817 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)))).
10820 (* constant 2175 *)
10821 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)))))).
10822 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)))))).
10825 (* constant 2176 *)
10826 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))))))))))).
10827 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)))))))).
10830 (* constant 2177 *)
10831 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)))))))).
10832 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))))))))).
10835 (* constant 2178 *)
10836 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)))))))))).
10837 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))))))))).
10840 (* constant 2179 *)
10841 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))))))))))).
10842 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))))))))).
10845 (* constant 2180 *)
10846 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))))))))).
10847 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)))))))).
10850 (* constant 2181 *)
10851 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)))))))).
10852 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)))))))).
10855 (* constant 2182 *)
10856 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)))))))))).
10857 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)))))))).
10860 (* constant 2183 *)
10861 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))))))))).
10862 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))))))))).
10865 (* constant 2184 *)
10866 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))))))))).
10867 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)))))))))).
10870 (* constant 2185 *)
10871 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)))))))))).
10872 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)))))))))).
10875 (* constant 2186 *)
10876 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)))))))))).
10877 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)))))))))).
10880 (* constant 2187 *)
10881 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)))))))))).
10882 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))))))))))))).
10885 (* constant 2188 *)
10886 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)))))))))).
10887 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)))))))))).
10890 (* constant 2189 *)
10891 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)))))))))).
10892 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)))))))))).
10895 (* constant 2190 *)
10896 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))))))))).
10897 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)))))))))).
10900 (* constant 2191 *)
10901 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))))))))).
10902 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)))))))))).
10905 (* constant 2192 *)
10906 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))))))))).
10907 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)))))))))).
10910 (* constant 2193 *)
10911 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)))))))))).
10912 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)))))))))).
10915 (* constant 2194 *)
10916 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))))))))).
10917 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)))))))))).
10920 (* constant 2195 *)
10921 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))))))))).
10922 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)))))))))).
10925 (* constant 2196 *)
10926 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))))))))).
10927 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)))))))))).
10930 (* constant 2197 *)
10931 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)))))))))).
10932 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)))))))))).
10935 (* constant 2198 *)
10936 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)))))))))).
10937 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)))))))))).
10940 (* constant 2199 *)
10941 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))))))))).
10942 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)))))))))).
10945 (* constant 2200 *)
10946 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))))))))))).
10947 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)))))))))).
10950 (* constant 2201 *)
10951 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))))))))))).
10952 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)))))))))))).
10955 (* constant 2202 *)
10956 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)))))))))).
10957 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))))))))))).
10960 (* constant 2203 *)
10961 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)))))))))).
10962 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)))))))))).
10965 (* constant 2204 *)
10966 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)))))))))).
10967 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)))))))))))).
10970 (* constant 2205 *)
10971 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)))))))))).
10972 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))))))))))).
10975 (* constant 2206 *)
10976 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)))))))))).
10977 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)))))))))))).
10980 (* constant 2207 *)
10981 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)))))))))).
10982 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)))))))))).
10985 (* constant 2208 *)
10986 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)))))))))).
10987 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)))))))))).
10990 (* constant 2209 *)
10991 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))))))))).
10992 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)))))))))).
10995 (* constant 2210 *)
10996 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)))))))).
10997 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))))))))).
11000 (* constant 2211 *)
11001 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)))))).
11002 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)))))))).
11005 (* constant 2212 *)
11006 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)))).
11007 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)))))).
11010 (* constant 2213 *)
11011 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))))).
11012 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)))).
11015 (* constant 2214 *)
11016 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)))))))).
11017 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)))))))).
11020 (* constant 2215 *)
11021 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))))))))).
11022 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)))))))).
11025 (* constant 2216 *)
11026 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)))))))).
11027 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))))))))).
11030 (* constant 2217 *)
11031 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)))))))).
11032 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))))))))).
11035 (* constant 2218 *)
11036 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))))))))))).
11037 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))))))))))).
11040 (* constant 2219 *)
11041 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)))))))).
11042 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))))))))).
11045 (* constant 2220 *)
11046 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)))))))).
11047 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))))))))).
11050 (* constant 2221 *)
11051 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)))))).
11052 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)))))))).
11055 (* constant 2222 *)
11056 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)))).
11057 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)))))).
11060 (* constant 2223 *)
11061 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)))))))))).
11062 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)))))))))).
11065 (* constant 2224 *)
11066 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))))))))).
11067 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)))))))))).
11070 (* constant 2225 *)
11071 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)))))))))).
11072 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))))))))).
11075 (* constant 2226 *)
11076 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))))))))).
11077 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)))))))))).
11080 (* constant 2227 *)
11081 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)))).
11082 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))))))))).
11085 (* constant 2228 *)
11086 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)).
11087 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)))).
11090 (* constant 2229 *)
11091 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))).
11092 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))).
11095 (* constant 2230 *)
11096 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)))))))))).
11097 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)))))))))).
11100 (* constant 2231 *)
11101 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)))))))))).
11102 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)))))))))).
11105 (* constant 2232 *)
11106 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)))))))))))))))).
11107 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)))))))))))))))).
11110 (* constant 2233 *)
11111 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)))))))))))))))).
11112 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)))))))))))))))))).
11115 (* constant 2234 *)
11116 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)))))))))))))))).
11117 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)))))))))))))))).
11120 (* constant 2235 *)
11121 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)))))))))))))))).
11122 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))))))))))))))))).
11125 (* constant 2236 *)
11126 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))))))))))))))))).
11127 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)))))))))))))))).
11130 (* constant 2237 *)
11131 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)))))))))))))))).
11132 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)))))))))))))))).
11135 (* constant 2238 *)
11136 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))))))))))).
11137 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))))))))))))))).
11140 (* constant 2239 *)
11141 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)))))))).
11142 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)))))))))).
11145 (* constant 2240 *)
11146 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))))).
11147 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))))))).
11150 (* constant 2241 *)
11151 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))))).
11152 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)))).
11155 (* constant 2242 *)
11156 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))))).
11157 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)))).
11160 (* constant 2243 *)
11161 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))))).
11162 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)))))).
11165 (* constant 2244 *)
11166 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))))).
11167 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))))).
11170 (* constant 2245 *)
11171 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))))).
11172 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)))).
11175 (* constant 2246 *)
11176 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))))).
11177 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))))).
11180 (* constant 2247 *)
11181 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))))))).
11182 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))))))).
11185 (* constant 2248 *)
11186 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))))))).
11187 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))))))).
11190 (* constant 2249 *)
11191 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))))))).
11192 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))))))).
11195 (* constant 2250 *)
11196 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))))))).
11197 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))))))).
11200 (* constant 2251 *)
11201 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)))).
11202 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))).
11205 (* constant 2252 *)
11206 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))))).
11207 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)))).
11210 (* constant 2253 *)
11211 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)))).
11212 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)))).
11215 (* constant 2254 *)
11216 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)))).
11217 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)))).
11220 (* constant 2255 *)
11221 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)))).
11222 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)))).
11225 (* constant 2256 *)
11226 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)))).
11227 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))))).
11230 (* constant 2257 *)
11231 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)))).
11232 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)))))).
11235 (* constant 2258 *)
11236 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)))).
11237 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))))).
11240 (* constant 2259 *)
11241 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))))))).
11242 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)))))).
11245 (* constant 2260 *)
11246 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))))))).
11247 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))))))).
11250 (* constant 2261 *)
11251 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))))))).
11252 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))))))).
11255 (* constant 2262 *)
11256 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))))))).
11257 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)))))))).
11260 (* constant 2263 *)
11261 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))))))).
11262 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))))))).
11265 (* constant 2264 *)
11266 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))))))).
11267 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))))))).
11270 (* constant 2265 *)
11271 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))))))).
11272 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)))))))).
11275 (* constant 2266 *)
11276 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))))))).
11277 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)))))))).
11280 (* constant 2267 *)
11281 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))))))))).
11282 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))))))))).
11285 (* constant 2268 *)
11286 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))))))))).
11287 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))))))))).
11290 (* constant 2269 *)
11291 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)))))))).
11292 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)))))))).
11295 (* constant 2270 *)
11296 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)))))))).
11297 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)))))))).
11300 (* constant 2271 *)
11301 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))))))).
11302 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))))))).
11305 (* constant 2272 *)
11306 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))))))).
11307 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)))))))).
11310 (* constant 2273 *)
11311 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))))).
11312 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)))))).
11315 (* constant 2274 *)
11316 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)))).
11317 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))))).
11320 (* constant 2275 *)
11321 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))))).
11322 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)))).
11325 (* constant 2276 *)
11326 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)))))).
11327 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))))).
11330 (* constant 2277 *)
11331 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)))))).
11332 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)))))).
11335 (* constant 2278 *)
11336 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)))))).
11337 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))))).
11340 (* constant 2279 *)
11341 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)))))).
11342 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)))))).
11345 (* constant 2280 *)
11346 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)))))).
11347 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))))).
11350 (* constant 2281 *)
11351 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)))))).
11352 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)))))).
11355 (* constant 2282 *)
11356 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)))))).
11357 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))))))).
11360 (* constant 2283 *)
11361 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))))).
11362 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))))))).
11365 (* constant 2284 *)
11366 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))))).
11367 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)))))).
11370 (* constant 2285 *)
11371 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))).
11372 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))))).
11375 (* constant 2286 *)
11376 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)).
11377 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))).
11380 (* constant 2287 *)
11381 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))))))))))).
11382 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)))))))))))).
11385 (* constant 2288 *)
11386 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))))))))).
11387 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)))))))))).
11390 (* constant 2289 *)
11391 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))))))))).
11392 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)))))))))).
11395 (* constant 2290 *)
11396 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)))))))))).
11397 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)))))))))).
11400 (* constant 2291 *)
11401 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))))))))).
11402 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)))))))))).
11405 (* constant 2292 *)
11406 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)))))))))).
11407 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)))))))))).
11410 (* constant 2293 *)
11411 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)))))).
11412 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)))))).
11415 (* constant 2294 *)
11416 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)))))))))).
11417 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)))))))))).
11420 (* constant 2295 *)
11421 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)))))))))).
11422 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)))))))))).
11425 (* constant 2296 *)
11426 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)))))))))).
11427 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)))))))))).
11430 (* constant 2297 *)
11431 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)))))))))).
11432 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))))))))))).
11435 (* constant 2298 *)
11436 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)))))))))))).
11437 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))))))))))).
11440 (* constant 2299 *)
11441 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)))))))))).
11442 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))))))))))).
11445 (* constant 2300 *)
11446 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)))))))).
11447 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)))))))))).
11450 (* constant 2301 *)
11451 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)))))).
11452 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)))))))).
11455 (* constant 2302 *)
11456 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))))))))).
11457 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))))))))).
11460 (* constant 2303 *)
11461 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)))))))))).
11462 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)))))))))))).
11465 (* constant 2304 *)
11466 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))))))))))))).
11467 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)))))))))))))).
11470 (* constant 2305 *)
11471 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))))))))))))).
11472 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)))))))))))))).
11475 (* constant 2306 *)
11476 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))))))))))))).
11477 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)))))))))))))).
11480 (* constant 2307 *)
11481 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))))))).
11482 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))))))))))))).
11485 (* constant 2308 *)
11486 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))))))).
11487 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)))))).
11490 (* constant 2309 *)
11491 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))))))))))))).
11492 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)))))))))))))).
11495 (* constant 2310 *)
11496 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))))))))))))).
11497 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))))))))))))).
11500 (* constant 2311 *)
11501 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))))))))))))).
11502 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)))))))))))))).
11505 (* constant 2312 *)
11506 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)))))))))))))).
11507 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)))))))))))))).
11510 (* constant 2313 *)
11511 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))))))))))))).
11512 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)))))))))))))).
11515 (* constant 2314 *)
11516 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))))))))))))))).
11517 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)))))))))))))).
11520 (* constant 2315 *)
11521 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)))))))))))))).
11522 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)))))))))))))).
11525 (* constant 2316 *)
11526 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)))))))).
11527 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))))))))))))).
11530 (* constant 2317 *)
11531 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)))))))))))))).
11532 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)))))))))))))).
11535 (* constant 2318 *)
11536 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)))))))))))))).
11537 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)))))))))))))).
11540 (* constant 2319 *)
11541 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))))))))))))))).
11542 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))))))))))))))).
11545 (* constant 2320 *)
11546 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))))))))))))))).
11547 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))))))))))))))).
11550 (* constant 2321 *)
11551 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)))))))))))))).
11552 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))))))))))))))).
11555 (* constant 2322 *)
11556 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))))))))))))))).
11557 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))))))))))))))))).
11560 (* constant 2323 *)
11561 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))))))))))))))).
11562 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))))))))))))))).
11565 (* constant 2324 *)
11566 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)))))))))))))))).
11567 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))))))))))))))).
11570 (* constant 2325 *)
11571 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))))))))))))).
11572 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)))))))))))))).
11575 (* constant 2326 *)
11576 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))))))).
11577 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))))))))))).
11580 (* constant 2327 *)
11581 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)))))))))))).
11582 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))))))))))))).
11585 (* constant 2328 *)
11586 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)))))))))).
11587 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))))))))))).
11590 (* constant 2329 *)
11591 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))))))).
11592 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))))))))).
11595 (* constant 2330 *)
11596 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)))).
11597 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)))))).
11600 (* constant 2331 *)
11601 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))).
11602 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)))).
11605 (* constant 2332 *)
11606 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)))))))))))))))).
11607 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)))))))))))))))).
11610 (* constant 2333 *)
11611 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)))))))))))))))).
11612 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))))))))))))))))).
11615 (* constant 2334 *)
11616 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)))))))))))))))).
11617 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))))))))))))))))))).
11620 (* constant 2335 *)
11621 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))))))))))))))))).
11622 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))))))))))))))))).
11625 (* constant 2336 *)
11626 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)))))))))))))))).
11627 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))))))))))))))))).
11630 (* constant 2337 *)
11631 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)))))))))))))))).
11632 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))))))))))))))))).
11635 (* constant 2338 *)
11636 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)))))))))).
11637 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)))))))))))))))).
11640 (* constant 2339 *)
11641 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))))).
11642 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)))))))))).
11645 (* constant 2340 *)
11646 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))))))))).
11647 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))))))))).
11650 (* constant 2341 *)
11651 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))))))))))))).
11652 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))))))))))))).
11655 (* constant 2342 *)
11656 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)))))))))))))).
11657 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)))))))))))))).
11660 (* constant 2343 *)
11661 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))))))))))))))).
11662 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))))))))))))))).
11665 (* constant 2344 *)
11666 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)))))))))))))).
11667 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)))))))))))))))).
11670 (* constant 2345 *)
11671 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)))))))))))))).
11672 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))))))))))))))))).
11675 (* constant 2346 *)
11676 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)))))))))))))))).
11677 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))))))))))))))).
11680 (* constant 2347 *)
11681 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))))))))))))))).
11682 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)))))))))))))))).
11685 (* constant 2348 *)
11686 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)))))))))))))))).
11687 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)))))))))))))))))).
11690 (* constant 2349 *)
11691 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)))))))))))))).
11692 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)))))))))))))))).
11695 (* constant 2350 *)
11696 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))))))))).
11697 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)))))))))))))).
11700 (* constant 2351 *)
11701 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)))))).
11702 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))))))))).
11705 (* constant 2352 *)
11706 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))))))))).
11707 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))))))))).
11710 (* constant 2353 *)
11711 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))))))))).
11712 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)))))))))).
11715 (* constant 2354 *)
11716 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)))))).
11717 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))))))))).
11720 (* constant 2355 *)
11721 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))))).
11722 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)))))).
11725 (* constant 2356 *)
11726 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))).
11727 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))))).
11730 (* constant 2357 *)
11731 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)))).
11732 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)))).
11735 (* constant 2358 *)
11736 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))).
11737 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)))))).
11740 (* constant 2359 *)
11741 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)))).
11742 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)))).
11745 (* constant 2360 *)
11746 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))).
11747 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)))).
11750 (* constant 2361 *)
11751 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))).
11752 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)))).
11755 (* constant 2362 *)
11756 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))))).
11757 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)))).
11760 (* constant 2363 *)
11761 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))).
11762 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)))).
11765 (* constant 2364 *)
11766 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)))).
11767 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)))).
11770 (* constant 2365 *)
11771 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)))))).
11772 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)))))).
11775 (* constant 2366 *)
11776 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)))))))).
11777 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)))))))).
11780 (* constant 2367 *)
11781 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))))))))).
11782 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))))))))).
11785 (* constant 2368 *)
11786 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))))))).
11787 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))))))).
11790 (* constant 2369 *)
11791 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))))))).
11792 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)))))).
11795 (* constant 2370 *)
11796 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))))).
11797 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))))))).
11800 (* constant 2371 *)
11801 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))).
11802 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))))).
11805 (* constant 2372 *)
11806 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)).
11807 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))).
11810 (* constant 2373 *)
11811 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)))))))).
11812 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)))))))).
11815 (* constant 2374 *)
11816 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))))))))).
11817 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))))))))).
11820 (* constant 2375 *)
11821 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)))))))).
11822 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))))))))).
11825 (* constant 2376 *)
11826 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))))))))).
11827 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))))))))).
11830 (* constant 2377 *)
11831 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)))))).
11832 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)))))).
11835 (* constant 2378 *)
11836 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)))))))))).
11837 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)))))))))).
11840 (* constant 2379 *)
11841 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)))))))))).
11842 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)))))))))).
11845 (* constant 2380 *)
11846 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))))))))))).
11847 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)))))))))).
11850 (* constant 2381 *)
11851 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)))))))))).
11852 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))))))))))).
11855 (* constant 2382 *)
11856 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)))))))).
11857 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)))))))))).
11860 (* constant 2383 *)
11861 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)))))).
11862 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)))))))).
11865 (* constant 2384 *)
11866 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))))))))))))).
11867 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))))))))))))).
11870 (* constant 2385 *)
11871 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))))))))))))).
11872 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))))))))))))).
11875 (* constant 2386 *)
11876 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)))))))))))))).
11877 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))))))))))))))).
11880 (* constant 2387 *)
11881 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)))))))))))))).
11882 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)))))))))))))).
11885 (* constant 2388 *)
11886 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))))))))).
11887 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))))))))))))).
11890 (* constant 2389 *)
11891 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)))))))).
11892 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))))))).
11895 (* constant 2390 *)
11896 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)).
11897 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)).
11900 (* constant 2391 *)
11901 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)).
11902 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))).
11905 (* constant 2392 *)
11906 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))).
11907 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))).
11910 (* constant 2393 *)
11911 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))).
11912 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))).
11915 (* constant 2394 *)
11916 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)))))))))))).
11917 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))))))))))).
11920 (* constant 2395 *)
11921 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))))))))))).
11922 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)))))))))))).
11925 (* constant 2396 *)
11926 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))))))))))).
11927 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))))))))))))).
11930 (* constant 2397 *)
11931 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)))))))))))).
11932 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)))))))))))).
11935 (* constant 2398 *)
11936 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)))))))))))).
11937 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)))))))))))).
11940 (* constant 2399 *)
11941 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))))))).
11942 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))))))))))).
11945 (* constant 2400 *)
11946 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))))))))))))).
11947 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)))))))))))))).
11950 (* constant 2401 *)
11951 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))))))))))))).
11952 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))))))))))))).
11955 (* constant 2402 *)
11956 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)))))))))))).
11957 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))))))))))))).
11960 (* constant 2403 *)
11961 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))))))))))))).
11962 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))))))))))))).
11965 (* constant 2404 *)
11966 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)))))))))))))).
11967 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))))))))))))).
11970 (* constant 2405 *)
11971 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))))))))))).
11972 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)))))))))))).
11975 (* constant 2406 *)
11976 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)))))).
11977 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))))))))).
11980 (* constant 2407 *)
11981 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))))))))))).
11982 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)))))))))))).
11985 (* constant 2408 *)
11986 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))))))))).
11987 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)))))))))).
11990 (* constant 2409 *)
11991 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))))))).
11992 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)))))))).
11995 (* constant 2410 *)
11996 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))))).
11997 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)))))).
12000 (* constant 2411 *)
12001 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))).
12002 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)))).
12005 (* constant 2412 *)
12006 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)).
12007 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))).
12010 (* constant 2413 *)
12011 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)))))))).
12012 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))))))))).
12015 (* constant 2414 *)
12016 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)))))))))).
12017 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)))))))).
12020 (* constant 2415 *)
12021 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)))))))).
12022 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))))))))).
12025 (* constant 2416 *)
12026 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))))))).
12027 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)))))).
12030 (* constant 2417 *)
12031 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)))))).
12032 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)))))).
12035 (* constant 2418 *)
12036 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))))).
12037 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)))).
12040 (* constant 2419 *)
12041 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))))).
12042 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)))).
12045 (* constant 2420 *)
12046 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))))))).
12047 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))))))).
12050 (* constant 2421 *)
12051 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)))))))))).
12052 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)))))))))).
12055 (* constant 2422 *)
12056 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))))))))).
12057 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)))))))))).
12060 (* constant 2423 *)
12061 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)))).
12062 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))))))))).
12065 (* constant 2424 *)
12066 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))).
12067 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)))).
12070 (* constant 2425 *)
12071 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))).
12072 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)).
12075 (* constant 2426 *)
12076 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))))))))))))))))).
12077 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)))))))))))))))).
12080 (* constant 2427 *)
12081 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)))))))))))))))).
12082 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))))))))))))))))).
12085 (* constant 2428 *)
12086 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))))))))))))))).
12087 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)))))))))))))))).
12090 (* constant 2429 *)
12091 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)))))))))).
12092 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))))))))))))))).
12095 (* constant 2430 *)
12096 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))))).
12097 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)))))))))).
12100 (* constant 2431 *)
12101 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)))))))))))))))).
12102 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)))))))))))))))).
12105 (* constant 2432 *)
12106 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)))))))))))))))).
12107 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))))))))))))))))).
12110 (* constant 2433 *)
12111 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))))))))))))))).
12112 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)))))))))))))))).
12115 (* constant 2434 *)
12116 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)))))))))).
12117 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))))))))))))))).
12120 (* constant 2435 *)
12121 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))))).
12122 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)))))))))).
12125 (* constant 2436 *)
12126 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))))).
12127 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))))).
12130 (* constant 2437 *)
12131 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))))).
12132 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))).
12135 (* constant 2438 *)
12136 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)))).
12137 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)))).
12140 (* constant 2439 *)
12141 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))))))))))))))))).
12142 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)))))))))))))))).
12145 (* constant 2440 *)
12146 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)))))))))))))))).
12147 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))))))))))))))))).
12150 (* constant 2441 *)
12151 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)))))))))))))))).
12152 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))))))))))))))))).
12155 (* constant 2442 *)
12156 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))))))))))))))).
12157 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)))))))))))))))).
12160 (* constant 2443 *)
12161 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)))))))))).
12162 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))))))))))))))).
12165 (* constant 2444 *)
12166 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))))).
12167 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)))))))))).
12170 (* constant 2445 *)
12171 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)))))))))))))))))))))).
12172 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))))))))))))))))))))).
12175 (* constant 2446 *)
12176 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)))))))))))))))))))).
12177 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)))))))))))))))))))).
12180 (* constant 2447 *)
12181 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))))))))))))))))))))).
12182 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))))))))))))))))))))).
12185 (* constant 2448 *)
12186 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)))))))))))))))))))))).
12187 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)))))))))))))))))))))).
12190 (* constant 2449 *)
12191 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)))))))))))))))))))))).
12192 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))))))))))))))))))))))).
12195 (* constant 2450 *)
12196 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)))))))))))))))))))))).
12197 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)))))))))))))))))))))).
12200 (* constant 2451 *)
12201 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))))))))))))))))))))).
12202 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))))))))))))))))))))).
12205 (* constant 2452 *)
12206 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)))))))))))))))))))))).
12207 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)))))))))))))))))))))).
12210 (* constant 2453 *)
12211 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)))))))))))))))))))))).
12212 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))))))))))))))))))))))).
12215 (* constant 2454 *)
12216 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)))))))))))))))))))))).
12217 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))))))))))))))))))))))).
12220 (* constant 2455 *)
12221 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))))))))))))))))))))).
12222 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))))))))))))))))))))).
12225 (* constant 2456 *)
12226 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))))))))))))))))))))).
12227 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))))))))))))))))))))).
12230 (* constant 2457 *)
12231 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))))))))))))))))))))).
12232 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))))))))))))))))))))).
12235 (* constant 2458 *)
12236 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))))))))))))))))))))).
12237 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)))))))))))))))))))))).
12240 (* constant 2459 *)
12241 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)))))))))))))))))))))).
12242 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))))))))))))))))))))))).
12245 (* constant 2460 *)
12246 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))))))))))))))))))))).
12247 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)))))))))))))))))))))).
12250 (* constant 2461 *)
12251 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))))))))))))))))))))).
12252 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)))))))))))))))))))))).
12255 (* constant 2462 *)
12256 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)))))))))))))))))))))).
12257 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)))))))))))))))))))))).
12260 (* constant 2463 *)
12261 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)))))))))))))))))))).
12262 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))))))))))))))))))))).
12265 (* constant 2464 *)
12266 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))))))))))))))).
12267 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)))))))))))))))))))).
12270 (* constant 2465 *)
12271 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)))))))))).
12272 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))))))))))))))).
12275 (* constant 2466 *)
12276 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))))).
12277 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)))))))))).
12280 (* constant 2467 *)
12281 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))))).
12282 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))))).
12285 (* constant 2468 *)
12286 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))))).
12287 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))))).
12290 (* constant 2469 *)
12291 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))))).
12292 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))).
12295 (* constant 2470 *)
12296 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)))).
12297 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)))).
12300 (* constant 2471 *)
12301 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))))).
12302 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)))).
12305 (* constant 2472 *)
12306 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)))).
12307 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)))).
12310 (* constant 2473 *)
12311 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)))))).
12312 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)))).
12315 (* constant 2474 *)
12316 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)))))).
12317 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))))).
12320 (* constant 2475 *)
12321 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))))).
12322 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)))))).
12325 (* constant 2476 *)
12326 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))))).
12327 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)))).
12330 (* constant 2477 *)
12331 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))))).
12332 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)))))).
12335 (* constant 2478 *)
12336 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))))).
12337 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))))).
12340 (* constant 2479 *)
12341 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))))).
12342 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)))).
12345 (* constant 2480 *)
12346 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))))).
12347 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))))).
12350 (* constant 2481 *)
12351 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))))))).
12352 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))))))).
12355 (* constant 2482 *)
12356 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))))))).
12357 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))))))).
12360 (* constant 2483 *)
12361 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))))))).
12362 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))))))).
12365 (* constant 2484 *)
12366 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))))))).
12367 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))))))).
12370 (* constant 2485 *)
12371 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)))).
12372 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))).
12375 (* constant 2486 *)
12376 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))))).
12377 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)))).
12380 (* constant 2487 *)
12381 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)))).
12382 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)))).
12385 (* constant 2488 *)
12386 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)))).
12387 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)))).
12390 (* constant 2489 *)
12391 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)))).
12392 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)))).
12395 (* constant 2490 *)
12396 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)))).
12397 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))))).
12400 (* constant 2491 *)
12401 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)))).
12402 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)))))).
12405 (* constant 2492 *)
12406 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)))).
12407 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))))).
12410 (* constant 2493 *)
12411 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))))))).
12412 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)))))).
12415 (* constant 2494 *)
12416 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))))))).
12417 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))))))).
12420 (* constant 2495 *)
12421 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))))))).
12422 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))))))).
12425 (* constant 2496 *)
12426 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))))))).
12427 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)))))))).
12430 (* constant 2497 *)
12431 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))))))).
12432 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))))))).
12435 (* constant 2498 *)
12436 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))))))).
12437 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))))))).
12440 (* constant 2499 *)
12441 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))))))).
12442 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)))))))).
12445 (* constant 2500 *)
12446 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))))))).
12447 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)))))))).
12450 (* constant 2501 *)
12451 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))))))))).
12452 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))))))))).
12455 (* constant 2502 *)
12456 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))))))))).
12457 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))))))))).
12460 (* constant 2503 *)
12461 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)))))))).
12462 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)))))))).
12465 (* constant 2504 *)
12466 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)))))))).
12467 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)))))))).
12470 (* constant 2505 *)
12471 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))))))).
12472 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))))))).
12475 (* constant 2506 *)
12476 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))))))).
12477 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)))))))).
12480 (* constant 2507 *)
12481 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).
12482 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)).
12485 (* constant 2508 *)
12486 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)).
12487 exact (fun (r0:l_e_st_eq_landau_n_rt_rat) => l_e_st_eq_landau_n_rt_satz90 r0).
12490 (* constant 2509 *)
12491 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)))).
12492 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))).
12495 (* constant 2510 *)
12496 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)).
12497 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)).
12500 (* constant 2511 *)
12501 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))).
12502 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)).
12505 (* constant 2512 *)
12506 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))).
12507 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))).
12510 (* constant 2513 *)
12511 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)))))).
12512 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))))))).
12515 (* constant 2514 *)
12516 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))))).
12517 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)))).
12520 (* constant 2515 *)
12521 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))))).
12522 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))))).
12525 (* constant 2516 *)
12526 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))))).
12527 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))))).
12530 (* constant 2517 *)
12531 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)))))).
12532 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))))))).
12535 (* constant 2518 *)
12536 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))))).
12537 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))))).
12540 (* constant 2519 *)
12541 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)))).
12542 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))))).
12545 (* constant 2520 *)
12546 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)).
12547 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))).
12550 (* constant 2521 *)
12551 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).
12552 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)).
12555 (* constant 2522 *)
12556 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))).
12557 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)))).
12560 (* constant 2523 *)
12561 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))).
12562 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)))).
12565 (* constant 2524 *)
12566 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)))).
12567 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))).
12570 (* constant 2525 *)
12571 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))).
12572 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)))).
12575 (* constant 2526 *)
12576 Definition l_e_st_eq_landau_n_rt_rp_1rp : l_e_st_eq_landau_n_rt_cut.
12577 exact (l_e_st_eq_landau_n_rt_rp_rpofrt l_e_st_eq_landau_n_rt_1rt).
12580 (* constant 2527 *)
12581 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)))))))).
12582 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)))))))).
12585 (* constant 2528 *)
12586 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)))))))).
12587 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)))))))))).
12590 (* constant 2529 *)
12591 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)))))))).
12592 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))))))))).
12595 (* constant 2530 *)
12596 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))).
12597 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)))))))).
12600 (* constant 2531 *)
12601 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)))))).
12602 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)))))).
12605 (* constant 2532 *)
12606 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)))))).
12607 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))))))).
12610 (* constant 2533 *)
12611 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))))))).
12612 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))))))).
12615 (* constant 2534 *)
12616 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)))))).
12617 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))))))).
12620 (* constant 2535 *)
12621 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)))))).
12622 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)))))))).
12625 (* constant 2536 *)
12626 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))).
12627 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)))))).
12630 (* constant 2537 *)
12631 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).
12632 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))).
12635 (* constant 2538 *)
12636 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)).
12637 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)).
12640 (* constant 2539 *)
12641 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).
12642 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)).
12645 (* constant 2540 *)
12646 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)).
12647 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)).
12650 (* constant 2541 *)
12651 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))).
12652 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)))).
12655 (* constant 2542 *)
12656 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))).
12657 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))))).
12660 (* constant 2543 *)
12661 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)).
12662 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))).
12665 (* constant 2544 *)
12666 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).
12667 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)).
12670 (* constant 2545 *)
12671 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)))))))).
12672 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)))))))).
12675 (* constant 2546 *)
12676 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))))))))).
12677 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))))))))).
12680 (* constant 2547 *)
12681 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)))))))).
12682 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)))))))).
12685 (* constant 2548 *)
12686 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)))))))).
12687 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))))))))).
12690 (* constant 2549 *)
12691 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))))))))).
12692 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))))))))).
12695 (* constant 2550 *)
12696 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))))).
12697 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))))).
12700 (* constant 2551 *)
12701 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))))))).
12702 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))))))).
12705 (* constant 2552 *)
12706 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)))))))).
12707 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))))))).
12710 (* constant 2553 *)
12711 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)))))))).
12712 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))))))).
12715 (* constant 2554 *)
12716 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))))))))).
12717 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))))))))).
12720 (* constant 2555 *)
12721 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))))))))).
12722 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))))))))).
12725 (* constant 2556 *)
12726 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))))))))).
12727 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)))))))))).
12730 (* constant 2557 *)
12731 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))))))).
12732 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))))))))).
12735 (* constant 2558 *)
12736 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))))).
12737 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))))))).
12740 (* constant 2559 *)
12741 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))).
12742 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))).
12745 (* constant 2560 *)
12746 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)))).
12747 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))).
12750 (* constant 2561 *)
12751 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)))).
12752 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)))).
12755 (* constant 2562 *)
12756 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)))).
12757 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)))))).
12760 (* constant 2563 *)
12761 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))))).
12762 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)))))).
12765 (* constant 2564 *)
12766 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))))).
12767 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))))).
12770 (* constant 2565 *)
12771 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))))).
12772 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))))).
12775 (* constant 2566 *)
12776 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)))))).
12777 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))))))).
12780 (* constant 2567 *)
12781 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)))))).
12782 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))))))).
12785 (* constant 2568 *)
12786 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)))))).
12787 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)))))).
12790 (* constant 2569 *)
12791 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)))).
12792 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))))))))))).
12795 (* constant 2570 *)
12796 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))))).
12797 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)))).
12800 (* constant 2571 *)
12801 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))))))))).
12802 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)))))))).
12805 (* constant 2572 *)
12806 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)))))))))).
12807 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))))))))).
12810 (* constant 2573 *)
12811 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)))))))))).
12812 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)))))))))).
12815 (* constant 2574 *)
12816 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))))))))).
12817 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))))))))).
12820 (* constant 2575 *)
12821 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))))))))).
12822 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))))))))).
12825 (* constant 2576 *)
12826 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))))))))).
12827 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))))))))).
12830 (* constant 2577 *)
12831 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)))))))).
12832 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)))))))).
12835 (* constant 2578 *)
12836 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)))))))))).
12837 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))))))))).
12840 (* constant 2579 *)
12841 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))))))))).
12842 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))))))))).
12845 (* constant 2580 *)
12846 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)))))).
12847 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))))))))))).
12850 (* constant 2581 *)
12851 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))))))))))).
12852 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))))))))).
12855 (* constant 2582 *)
12856 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))))))))))).
12857 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))))))))))).
12860 (* constant 2583 *)
12861 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))))))))))).
12862 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)))))))))))).
12865 (* constant 2584 *)
12866 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)))))))))))).
12867 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))))))))))))).
12870 (* constant 2585 *)
12871 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))))))))))).
12872 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))))))))))).
12875 (* constant 2586 *)
12876 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))))))))))))).
12877 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)))))))))))).
12880 (* constant 2587 *)
12881 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))))))))))))).
12882 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))))))))))))).
12885 (* constant 2588 *)
12886 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)))))))))))).
12887 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)))))))))))).
12890 (* constant 2589 *)
12891 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)))))))))))).
12892 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)))))))))))).
12895 (* constant 2590 *)
12896 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))))))))))).
12897 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))))))))))))).
12900 (* constant 2591 *)
12901 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)))))))))))).
12902 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)))))))))))).
12905 (* constant 2592 *)
12906 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))))))))))))).
12907 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)))))))))))).
12910 (* constant 2593 *)
12911 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))))))))))).
12912 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))))))))))).
12915 (* constant 2594 *)
12916 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))))).
12917 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))))))))).
12920 (* constant 2595 *)
12921 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)))))).
12922 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))))))).
12925 (* constant 2596 *)
12926 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)))).
12927 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))))).
12930 (* constant 2597 *)
12931 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)).
12932 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))).
12935 (* constant 2598 *)
12936 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).
12937 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)).
12940 (* constant 2599 *)
12941 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))))))))))))).
12942 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)))))))))))).
12945 (* constant 2600 *)
12946 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))))))))))).
12947 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))))))))))).
12950 (* constant 2601 *)
12951 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))))))))))).
12952 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))))))))))))).
12955 (* constant 2602 *)
12956 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))))))))))).
12957 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)))))))))))).
12960 (* constant 2603 *)
12961 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))))))))).
12962 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)))))))).
12965 (* constant 2604 *)
12966 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)))))))).
12967 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)))))))))))))).
12970 (* constant 2605 *)
12971 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))).
12972 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)))))))).
12975 (* constant 2606 *)
12976 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))).
12977 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))).
12980 (* constant 2607 *)
12981 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))).
12982 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)))).
12985 (* constant 2608 *)
12986 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))))))))).
12987 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))))))))).
12990 (* constant 2609 *)
12991 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)))))))))).
12992 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)))))))))).
12995 (* constant 2610 *)
12996 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))))))))))).
12997 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))))))))))).
13000 (* constant 2611 *)
13001 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))))))))))).
13002 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))))))))))).
13005 (* constant 2612 *)
13006 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))))))))))).
13007 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)))))))))))).
13010 (* constant 2613 *)
13011 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)))))))))))).
13012 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))))))))))).
13015 (* constant 2614 *)
13016 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))))))))))).
13017 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))))))))))).
13020 (* constant 2615 *)
13021 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))))))))))))).
13022 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))))))))))).
13025 (* constant 2616 *)
13026 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)))))))))).
13027 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))))))))))).
13030 (* constant 2617 *)
13031 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)))))))))).
13032 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))))))))))).
13035 (* constant 2618 *)
13036 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))))))))))).
13037 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)))))))))))).
13040 (* constant 2619 *)
13041 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))))))))))).
13042 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))))))))))).
13045 (* constant 2620 *)
13046 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)))))))))).
13047 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)))))))))).
13050 (* constant 2621 *)
13051 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))))))))))))).
13052 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))))))))))).
13055 (* constant 2622 *)
13056 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))))))))))).
13057 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))))))))))))).
13060 (* constant 2623 *)
13061 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)))))))))))).
13062 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))))))))))).
13065 (* constant 2624 *)
13066 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)))))))))).
13067 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))))))))))).
13070 (* constant 2625 *)
13071 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))))).
13072 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)))))))))).
13075 (* constant 2626 *)
13076 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))).
13077 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))))).
13080 (* constant 2627 *)
13081 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).
13082 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))).
13085 (* constant 2628 *)
13086 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)).
13087 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)).
13090 (* constant 2629 *)
13091 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)))))).
13092 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))))).
13095 (* constant 2630 *)
13096 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)))))).
13097 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)))))).
13100 (* constant 2631 *)
13101 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)))))).
13102 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))))).
13105 (* constant 2632 *)
13106 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)))))).
13107 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)))))).
13110 (* constant 2633 *)
13111 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)))))).
13112 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))))).
13115 (* constant 2634 *)
13116 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)))))).
13117 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)))))).
13120 (* constant 2635 *)
13121 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)))))).
13122 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))))))).
13125 (* constant 2636 *)
13126 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)))).
13127 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)))).
13130 (* constant 2637 *)
13131 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)))).
13132 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))))).
13135 (* constant 2638 *)
13136 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))))).
13137 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))))).
13140 (* constant 2639 *)
13141 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))).
13142 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)))).
13145 (* constant 2640 *)
13146 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))).
13147 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)))))).
13150 (* constant 2641 *)
13151 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))).
13152 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))).
13155 (* constant 2642 *)
13156 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)).
13157 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))).
13160 (* constant 2643 *)
13161 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)).
13162 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))).
13165 (* constant 2644 *)
13166 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)))).
13167 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))).
13170 (* constant 2645 *)
13171 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)).
13172 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))).
13175 (* constant 2646 *)
13176 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))).
13177 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))).
13180 (* constant 2647 *)
13181 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))))).
13182 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))))).
13185 (* constant 2648 *)
13186 Definition l_e_st_eq_landau_n_rt_rp_ratrp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), Prop).
13187 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).
13190 (* constant 2649 *)
13191 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)).
13192 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).
13195 (* constant 2650 *)
13196 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).
13197 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)).
13200 (* constant 2651 *)
13201 Definition l_e_st_eq_landau_n_rt_rp_natrp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), Prop).
13202 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).
13205 (* constant 2652 *)
13206 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)).
13207 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).
13210 (* constant 2653 *)
13211 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)))).
13212 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)))).
13215 (* constant 2654 *)
13216 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)).
13217 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)))).
13220 (* constant 2655 *)
13221 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))).
13222 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)))).
13225 (* constant 2656 *)
13226 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))).
13227 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))))).
13230 (* constant 2657 *)
13231 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)))).
13232 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)))).
13235 (* constant 2658 *)
13236 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)))).
13237 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)))).
13240 (* constant 2659 *)
13241 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)))).
13242 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))).
13245 (* constant 2660 *)
13246 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)))).
13247 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))))).
13250 (* constant 2661 *)
13251 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))).
13252 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)).
13255 (* constant 2662 *)
13256 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)))).
13257 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))).
13260 (* constant 2663 *)
13261 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))).
13262 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))).
13265 (* constant 2664 *)
13266 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))).
13267 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))).
13270 (* constant 2665 *)
13271 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))).
13272 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))).
13275 (* constant 2666 *)
13276 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).
13277 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))).
13280 (* constant 2667 *)
13281 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)))).
13282 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))).
13285 (* constant 2668 *)
13286 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))).
13287 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))).
13290 (* constant 2669 *)
13291 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)).
13292 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)).
13295 (* constant 2670 *)
13296 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)))))).
13297 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))))).
13300 (* constant 2671 *)
13301 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))))).
13302 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))))).
13305 (* constant 2672 *)
13306 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))).
13307 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).
13310 (* constant 2673 *)
13311 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).
13312 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).
13315 (* constant 2674 *)
13316 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)))).
13317 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)).
13320 (* constant 2675 *)
13321 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)).
13322 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)).
13325 (* constant 2676 *)
13326 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)))).
13327 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))).
13330 (* constant 2677 *)
13331 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))).
13332 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)))).
13335 (* constant 2678 *)
13336 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).
13337 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))).
13340 (* constant 2679 *)
13341 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)).
13342 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)).
13345 (* constant 2680 *)
13346 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)))))).
13347 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))))).
13350 (* constant 2681 *)
13351 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))))).
13352 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))))).
13355 (* constant 2682 *)
13356 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))).
13357 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).
13360 (* constant 2683 *)
13361 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).
13362 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).
13365 (* constant 2684 *)
13366 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)))).
13367 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)).
13370 (* constant 2685 *)
13371 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)).
13372 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)).
13375 (* constant 2686 *)
13376 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))))))))).
13377 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))))))))).
13380 (* constant 2687 *)
13381 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))))))))).
13382 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))))))))).
13385 (* constant 2688 *)
13386 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)))))))))).
13387 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)))))))))).
13390 (* constant 2689 *)
13391 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)))))))))).
13392 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)))))))))).
13395 (* constant 2690 *)
13396 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))))))))).
13397 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)))))))))).
13400 (* constant 2691 *)
13401 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)))).
13402 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))))))))).
13405 (* constant 2692 *)
13406 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))))).
13407 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)))).
13410 (* constant 2693 *)
13411 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)))).
13412 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))))).
13415 (* constant 2694 *)
13416 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)))))).
13417 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))))).
13420 (* constant 2695 *)
13421 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)))).
13422 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))))).
13425 (* constant 2696 *)
13426 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))))))).
13427 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)))))).
13430 (* constant 2697 *)
13431 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))).
13432 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)))).
13435 (* constant 2698 *)
13436 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)))).
13437 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)))).
13440 (* constant 2699 *)
13441 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)))).
13442 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))))).
13445 (* constant 2700 *)
13446 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)))).
13447 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)))).
13450 (* constant 2701 *)
13451 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)))).
13452 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))).
13455 (* constant 2702 *)
13456 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))))).
13457 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)))).
13460 (* constant 2703 *)
13461 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)))).
13462 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)))).
13465 (* constant 2704 *)
13466 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))))).
13467 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)))).
13470 (* constant 2705 *)
13471 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))))))))).
13472 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))))))))).
13475 (* constant 2706 *)
13476 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))))))))).
13477 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))))))))).
13480 (* constant 2707 *)
13481 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)))))))))).
13482 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)))))))))).
13485 (* constant 2708 *)
13486 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)))))))))).
13487 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)))))))))).
13490 (* constant 2709 *)
13491 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))))))))).
13492 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)))))))))).
13495 (* constant 2710 *)
13496 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)))).
13497 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))))))))).
13500 (* constant 2711 *)
13501 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))))).
13502 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)))).
13505 (* constant 2712 *)
13506 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)))))).
13507 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)))))).
13510 (* constant 2713 *)
13511 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))))))).
13512 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)))))).
13515 (* constant 2714 *)
13516 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))))))).
13517 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))))))).
13520 (* constant 2715 *)
13521 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)))))).
13522 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))))))).
13525 (* constant 2716 *)
13526 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))))))).
13527 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))))))).
13530 (* constant 2717 *)
13531 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)))))).
13532 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))))))).
13535 (* constant 2718 *)
13536 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))))))))).
13537 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)))))))).
13540 (* constant 2719 *)
13541 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)))))).
13542 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))))))).
13545 (* constant 2720 *)
13546 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)))).
13547 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)))))).
13550 (* constant 2721 *)
13551 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)))).
13552 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)))).
13555 (* constant 2722 *)
13556 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))).
13557 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)).
13560 (* constant 2723 *)
13561 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)))).
13562 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))).
13565 (* constant 2724 *)
13566 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))).
13567 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))).
13570 (* constant 2725 *)
13571 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)))).
13572 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))).
13575 (* constant 2726 *)
13576 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)))).
13577 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)))).
13580 (* constant 2727 *)
13581 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)))).
13582 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)))).
13585 (* constant 2728 *)
13586 Definition l_e_st_eq_landau_n_rt_rp_nt_natt : Type.
13587 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)).
13590 (* constant 2729 *)
13591 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)).
13592 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)).
13595 (* constant 2730 *)
13596 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)).
13597 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)).
13600 (* constant 2731 *)
13601 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)).
13602 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))).
13605 (* constant 2732 *)
13606 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).
13607 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).
13610 (* constant 2733 *)
13611 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).
13612 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).
13615 (* constant 2734 *)
13616 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).
13617 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).
13620 (* constant 2735 *)
13621 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)).
13622 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)).
13625 (* constant 2736 *)
13626 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).
13627 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).
13630 (* constant 2737 *)
13631 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)).
13632 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).
13635 (* constant 2738 *)
13636 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)))))).
13637 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))))).
13640 (* constant 2739 *)
13641 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))))).
13642 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))))).
13645 (* constant 2740 *)
13646 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)))).
13647 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))).
13650 (* constant 2741 *)
13651 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))).
13652 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))).
13655 (* constant 2742 *)
13656 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)))).
13657 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)).
13660 (* constant 2743 *)
13661 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))).
13662 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).
13665 (* constant 2744 *)
13666 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).
13667 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)).
13670 (* constant 2745 *)
13671 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)))).
13672 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)))).
13675 (* constant 2746 *)
13676 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))).
13677 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)))).
13680 (* constant 2747 *)
13681 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).
13682 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)).
13685 (* constant 2748 *)
13686 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)))).
13687 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)))).
13690 (* constant 2749 *)
13691 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))).
13692 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)))).
13695 (* constant 2750 *)
13696 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))).
13697 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)).
13700 (* constant 2751 *)
13701 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))).
13702 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)).
13705 (* constant 2752 *)
13706 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))).
13707 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)).
13710 (* constant 2753 *)
13711 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))).
13712 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)).
13715 (* constant 2754 *)
13716 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))).
13717 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)).
13720 (* constant 2755 *)
13721 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))).
13722 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)).
13725 (* constant 2756 *)
13726 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).
13727 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)).
13730 (* constant 2757 *)
13731 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).
13732 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)).
13735 (* constant 2758 *)
13736 Definition l_e_st_eq_landau_n_rt_rp_nt_1t : l_e_st_eq_landau_n_rt_rp_nt_natt.
13737 exact (l_e_st_eq_landau_n_rt_rp_nt_nttofnt l_e_st_eq_landau_n_1).
13740 (* constant 2759 *)
13741 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).
13742 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))).
13745 (* constant 2760 *)
13746 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)).
13747 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)).
13750 (* constant 2761 *)
13751 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).
13752 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)).
13755 (* constant 2762 *)
13756 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))))).
13757 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))).
13760 (* constant 2763 *)
13761 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))).
13762 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))))).
13765 (* constant 2764 *)
13766 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).
13767 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).
13770 (* constant 2765 *)
13771 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).
13772 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))).
13775 (* constant 2766 *)
13776 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)))).
13777 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)))).
13780 (* constant 2767 *)
13781 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))))).
13782 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))))).
13785 (* constant 2768 *)
13786 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)))))).
13787 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)))))).
13790 (* constant 2769 *)
13791 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)))).
13792 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))))).
13795 (* constant 2770 *)
13796 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)))).
13797 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))))).
13800 (* constant 2771 *)
13801 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).
13802 exact (fun (x:l_e_st_eq_landau_n_rt_rp_nt_natt) => l_e_st_eq_landau_n_rt_rp_nt_satz156a x).
13805 (* constant 2772 *)
13806 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))).
13807 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))).
13810 (* constant 2773 *)
13811 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)))).
13812 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)))).
13815 (* constant 2774 *)
13816 Definition l_e_st_eq_landau_n_rt_rp_rtt_ratt : Type.
13817 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)).
13820 (* constant 2775 *)
13821 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)).
13822 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)).
13825 (* constant 2776 *)
13826 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)).
13827 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)).
13830 (* constant 2777 *)
13831 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)).
13832 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))).
13835 (* constant 2778 *)
13836 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).
13837 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).
13840 (* constant 2779 *)
13841 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).
13842 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).
13845 (* constant 2780 *)
13846 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).
13847 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).
13850 (* constant 2781 *)
13851 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).
13852 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).
13855 (* constant 2782 *)
13856 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)).
13857 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).
13860 (* constant 2783 *)
13861 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)))))).
13862 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))))).
13865 (* constant 2784 *)
13866 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))))).
13867 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))))).
13870 (* constant 2785 *)
13871 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)))).
13872 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))).
13875 (* constant 2786 *)
13876 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))).
13877 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))).
13880 (* constant 2787 *)
13881 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)))).
13882 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)).
13885 (* constant 2788 *)
13886 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))).
13887 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).
13890 (* constant 2789 *)
13891 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).
13892 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)).
13895 (* constant 2790 *)
13896 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)))).
13897 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)))).
13900 (* constant 2791 *)
13901 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))).
13902 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)))).
13905 (* constant 2792 *)
13906 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).
13907 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)).
13910 (* constant 2793 *)
13911 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)))).
13912 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)))).
13915 (* constant 2794 *)
13916 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))).
13917 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)))).
13920 (* constant 2795 *)
13921 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))).
13922 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)).
13925 (* constant 2796 *)
13926 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))).
13927 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)).
13930 (* constant 2797 *)
13931 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))).
13932 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)).
13935 (* constant 2798 *)
13936 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))).
13937 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)).
13940 (* constant 2799 *)
13941 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))).
13942 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)).
13945 (* constant 2800 *)
13946 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))).
13947 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)).
13950 (* constant 2801 *)
13951 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).
13952 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).
13955 (* constant 2802 *)
13956 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)).
13957 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)).
13960 (* constant 2803 *)
13961 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).
13962 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)).
13965 (* constant 2804 *)
13966 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)))).
13967 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)))).
13970 (* constant 2805 *)
13971 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))))).
13972 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)))))).
13975 (* constant 2806 *)
13976 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))))).
13977 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)))))).
13980 (* constant 2807 *)
13981 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))))).
13982 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))))).
13985 (* constant 2808 *)
13986 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)))).
13987 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))))).
13990 (* constant 2809 *)
13991 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))).
13992 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)))).
13995 (* constant 2810 *)
13996 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))).
13997 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))))).
14000 (* constant 2811 *)
14001 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))).
14002 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))).
14005 (* constant 2812 *)
14006 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))).
14007 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))).
14010 (* constant 2813 *)
14011 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))).
14012 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))).
14015 (* constant 2814 *)
14016 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))).
14017 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)).
14020 (* constant 2815 *)
14021 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))).
14022 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))).
14025 (* constant 2816 *)
14026 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))).
14027 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))).
14030 (* constant 2817 *)
14031 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)))).
14032 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))).
14035 (* constant 2818 *)
14036 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))).
14037 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)))).
14040 (* constant 2819 *)
14041 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))))).
14042 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)))))).
14045 (* constant 2820 *)
14046 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))))).
14047 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)))))).
14050 (* constant 2821 *)
14051 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)))))).
14052 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))))).
14055 (* constant 2822 *)
14056 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))))).
14057 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))))))).
14060 (* constant 2823 *)
14061 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))))).
14062 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)))))).
14065 (* constant 2824 *)
14066 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))))).
14067 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))))).
14070 (* constant 2825 *)
14071 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)))).
14072 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)))).
14075 (* constant 2826 *)
14076 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)))).
14077 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))))).
14080 (* constant 2827 *)
14081 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)).
14082 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)))).
14085 (* constant 2828 *)
14086 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)).
14087 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)).
14090 (* constant 2829 *)
14091 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))).
14092 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))))).
14095 (* constant 2830 *)
14096 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)))).
14097 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))).
14100 (* constant 2831 *)
14101 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))).
14102 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)))).
14105 (* constant 2832 *)
14106 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))).
14107 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)))).
14110 (* constant 2833 *)
14111 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)))).
14112 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))))).
14115 (* constant 2834 *)
14116 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)))).
14117 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))))).
14120 (* constant 2835 *)
14121 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))))).
14122 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)))).
14125 (* constant 2836 *)
14126 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))))).
14127 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))))).
14130 (* constant 2837 *)
14131 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))))))).
14132 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))))).
14135 (* constant 2838 *)
14136 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))))))).
14137 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)))))).
14140 (* constant 2839 *)
14141 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)))))).
14142 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))))))).
14145 (* constant 2840 *)
14146 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))))))).
14147 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)))))).
14150 (* constant 2841 *)
14151 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)))))).
14152 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)))))))).
14155 (* constant 2842 *)
14156 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)))))).
14157 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))))))).
14160 (* constant 2843 *)
14161 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))))))).
14162 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))))))).
14165 (* constant 2844 *)
14166 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)))))).
14167 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))))))).
14170 (* constant 2845 *)
14171 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)))).
14172 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)))))).
14175 (* constant 2846 *)
14176 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)))).
14177 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))))).
14180 (* constant 2847 *)
14181 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))).
14182 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)))).
14185 (* constant 2848 *)
14186 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)))).
14187 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))).
14190 (* constant 2849 *)
14191 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)))).
14192 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)))).
14195 (* constant 2850 *)
14196 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))).
14197 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)))).
14200 (* constant 2851 *)
14201 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)))).
14202 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))).
14205 (* constant 2852 *)
14206 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))).
14207 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)))).
14210 (* constant 2853 *)
14211 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)))).
14212 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)))).
14215 (* constant 2854 *)
14216 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))))))).
14217 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))))))).
14220 (* constant 2855 *)
14221 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)))))))))).
14222 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)))))))))).
14225 (* constant 2856 *)
14226 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)))))))))).
14227 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)))))))))).
14230 (* constant 2857 *)
14231 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))))))))))).
14232 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)))))))))).
14235 (* constant 2858 *)
14236 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)))))))).
14237 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))))))))).
14240 (* constant 2859 *)
14241 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))))).
14242 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)))))).
14245 (* constant 2860 *)
14246 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)))))).
14247 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))))).
14250 (* constant 2861 *)
14251 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))))))).
14252 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)))))).
14255 (* constant 2862 *)
14256 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))))).
14257 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))))).
14260 (* constant 2863 *)
14261 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)))))).
14262 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)))))).
14265 (* constant 2864 *)
14266 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)))))))).
14267 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))))))).
14270 (* constant 2865 *)
14271 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))))))).
14272 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))))))).
14275 (* constant 2866 *)
14276 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))))))).
14277 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)))))))).
14280 (* constant 2867 *)
14281 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))))).
14282 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))))))).
14285 (* constant 2868 *)
14286 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)))).
14287 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)))).
14290 (* constant 2869 *)
14291 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)))).
14292 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)))).
14295 (* constant 2870 *)
14296 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)))).
14297 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)))).
14300 (* constant 2871 *)
14301 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)))).
14302 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))))).
14305 (* constant 2872 *)
14306 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)))).
14307 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)))).
14310 (* constant 2873 *)
14311 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)))))).
14312 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))))).
14315 (* constant 2874 *)
14316 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)))))).
14317 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)))))).
14320 (* constant 2875 *)
14321 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))))).
14322 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)))))).
14325 (* constant 2876 *)
14326 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))))).
14327 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))))).
14330 (* constant 2877 *)
14331 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))))).
14332 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)))))).
14335 (* constant 2878 *)
14336 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)))))).
14337 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))))))).
14340 (* constant 2879 *)
14341 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)))).
14342 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))))).
14345 (* constant 2880 *)
14346 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))))).
14347 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))))).
14350 (* constant 2881 *)
14351 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))))).
14352 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))))).
14355 (* constant 2882 *)
14356 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)))))))).
14357 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)))))))).
14360 (* constant 2883 *)
14361 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))))))))))))).
14362 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))))))))))).
14365 (* constant 2884 *)
14366 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))))))))))))).
14367 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))))))))))))).
14370 (* constant 2885 *)
14371 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)))))))))))).
14372 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))))))))))))).
14375 (* constant 2886 *)
14376 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))))))))))))).
14377 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))))))))))).
14380 (* constant 2887 *)
14381 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))))))))))))).
14382 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))))))))))).
14385 (* constant 2888 *)
14386 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))))))))))).
14387 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)))))))))))).
14390 (* constant 2889 *)
14391 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)))))))))))).
14392 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)))))))))))).
14395 (* constant 2890 *)
14396 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))))))))))))).
14397 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)))))))))))).
14400 (* constant 2891 *)
14401 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))))))))))))).
14402 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))))))))))).
14405 (* constant 2892 *)
14406 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))))))))))).
14407 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)))))))))))))).
14410 (* constant 2893 *)
14411 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)))))))))))).
14412 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))))))))))).
14415 (* constant 2894 *)
14416 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)))))))))))).
14417 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))))))))))).
14420 (* constant 2895 *)
14421 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))))))))))).
14422 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))))))))))).
14425 (* constant 2896 *)
14426 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)))))))))).
14427 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)))))))))))).
14430 (* constant 2897 *)
14431 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)))))))))).
14432 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)))))))))).
14435 (* constant 2898 *)
14436 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)))))))))).
14437 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))))))))))).
14440 (* constant 2899 *)
14441 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)))))))))).
14442 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)))))))))).
14445 (* constant 2900 *)
14446 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)))))))))).
14447 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))))))))))).
14450 (* constant 2901 *)
14451 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)))))))))).
14452 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)))))))))).
14455 (* constant 2902 *)
14456 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))))))))))).
14457 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)))))))))))).
14460 (* constant 2903 *)
14461 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)))))))))).
14462 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))))))))))).
14465 (* constant 2904 *)
14466 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)))))))))).
14467 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))))))))))).
14470 (* constant 2905 *)
14471 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)))))))))).
14472 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)))))))))).
14475 (* constant 2906 *)
14476 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)))).
14477 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)))).
14480 (* constant 2907 *)
14481 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))))).
14482 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))))).
14485 (* constant 2908 *)
14486 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))))).
14487 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)))))).
14490 (* constant 2909 *)
14491 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))).
14492 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))))).
14495 (* constant 2910 *)
14496 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))))))))))).
14497 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))))))))))).
14500 (* constant 2911 *)
14501 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))))))))))).
14502 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))))))))))).
14505 (* constant 2912 *)
14506 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)))))))))).
14507 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))))))))))).
14510 (* constant 2913 *)
14511 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))))))).
14512 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)))))))))).
14515 (* constant 2914 *)
14516 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))))))).
14517 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))))))).
14520 (* constant 2915 *)
14521 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))))))).
14522 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))))))).
14525 (* constant 2916 *)
14526 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))))))))).
14527 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))))))))).
14530 (* constant 2917 *)
14531 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)))))))))).
14532 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)))))))))).
14535 (* constant 2918 *)
14536 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)))))))))).
14537 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)))))))))).
14540 (* constant 2919 *)
14541 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)))))))))).
14542 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)))))))))).
14545 (* constant 2920 *)
14546 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)))))))))).
14547 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))))))))))).
14550 (* constant 2921 *)
14551 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)))))))).
14552 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)))))))))).
14555 (* constant 2922 *)
14556 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)))))).
14557 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)))))))).
14560 (* constant 2923 *)
14561 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)).
14562 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)).
14565 (* constant 2924 *)
14566 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)).
14567 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)).
14570 (* constant 2925 *)
14571 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))).
14572 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))).
14575 (* constant 2926 *)
14576 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))))).
14577 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)))).
14580 (* constant 2927 *)
14581 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))))).
14582 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)))))).
14585 (* constant 2928 *)
14586 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))))).
14587 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))))).
14590 (* constant 2929 *)
14591 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))))).
14592 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)))))).
14595 (* constant 2930 *)
14596 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))))).
14597 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))))))).
14600 (* constant 2931 *)
14601 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)))).
14602 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))))).
14605 (* constant 2932 *)
14606 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)))).
14607 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))))).
14610 (* constant 2933 *)
14611 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))))).
14612 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)))).
14615 (* constant 2934 *)
14616 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))))).
14617 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)))))).
14620 (* constant 2935 *)
14621 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))))).
14622 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)))))).
14625 (* constant 2936 *)
14626 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))))).
14627 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)))))).
14630 (* constant 2937 *)
14631 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))))).
14632 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))))))).
14635 (* constant 2938 *)
14636 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)))).
14637 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))))).
14640 (* constant 2939 *)
14641 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)))).
14642 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))))).
14645 (* constant 2940 *)
14646 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))))).
14647 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)))).
14650 (* constant 2941 *)
14651 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))).
14652 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))).
14655 (* constant 2942 *)
14656 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))).
14657 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))).
14660 (* constant 2943 *)
14661 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))))).
14662 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))))).
14665 (* constant 2944 *)
14666 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)))))).
14667 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))))).
14670 (* constant 2945 *)
14671 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)))))).
14672 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))))))).
14675 (* constant 2946 *)
14676 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)))))).
14677 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)))))))).
14680 (* constant 2947 *)
14681 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))))).
14682 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)))))).
14685 (* constant 2948 *)
14686 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)).
14687 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))))).
14690 (* constant 2949 *)
14691 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).
14692 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)).
14695 (* constant 2950 *)
14696 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)).
14697 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)).
14700 (* constant 2951 *)
14701 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))).
14702 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))).
14705 (* constant 2952 *)
14706 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))).
14707 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))).
14710 (* constant 2953 *)
14711 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))).
14712 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)))).
14715 (* constant 2954 *)
14716 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))).
14717 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)))).
14720 (* constant 2955 *)
14721 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)))).
14722 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)))).
14725 (* constant 2956 *)
14726 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))).
14727 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))).
14730 (* constant 2957 *)
14731 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))).
14732 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))).
14735 (* constant 2958 *)
14736 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))).
14737 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)))).
14740 (* constant 2959 *)
14741 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))).
14742 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)))).
14745 (* constant 2960 *)
14746 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)))).
14747 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)))).
14750 (* constant 2961 *)
14751 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))))).
14752 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)))).
14755 (* constant 2962 *)
14756 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)))).
14757 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)))).
14760 (* constant 2963 *)
14761 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))).
14762 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))).
14765 (* constant 2964 *)
14766 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)))))).
14767 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))))).
14770 (* constant 2965 *)
14771 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))))).
14772 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)))))).
14775 (* constant 2966 *)
14776 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)))))).
14777 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)))))).
14780 (* constant 2967 *)
14781 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))))).
14782 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)))).
14785 (* constant 2968 *)
14786 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))).
14787 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)))).
14790 (* constant 2969 *)
14791 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))).
14792 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)))).
14795 (* constant 2970 *)
14796 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))).
14797 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)))).
14800 (* constant 2971 *)
14801 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)))).
14802 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)))).
14805 (* constant 2972 *)
14806 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))))).
14807 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))))).
14810 (* constant 2973 *)
14811 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)))))).
14812 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))))).
14815 (* constant 2974 *)
14816 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))))).
14817 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))))).
14820 (* constant 2975 *)
14821 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)))))))).
14822 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))))))).
14825 (* constant 2976 *)
14826 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)))))))).
14827 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)))))).
14830 (* constant 2977 *)
14831 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)))))))).
14832 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))))))))).
14835 (* constant 2978 *)
14836 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))))))).
14837 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))))))).
14840 (* constant 2979 *)
14841 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)))))))).
14842 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)))))).
14845 (* constant 2980 *)
14846 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)))))))).
14847 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)))))))).
14850 (* constant 2981 *)
14851 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)))))))).
14852 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)))))).
14855 (* constant 2982 *)
14856 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)))))).
14857 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))))))).
14860 (* constant 2983 *)
14861 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))))))).
14862 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)))))).
14865 (* constant 2984 *)
14866 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))))).
14867 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)))))).
14870 (* constant 2985 *)
14871 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))))).
14872 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)))))).
14875 (* constant 2986 *)
14876 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)))))).
14877 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)))))).
14880 (* constant 2987 *)
14881 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)))))).
14882 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)))))).
14885 (* constant 2988 *)
14886 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))))))).
14887 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)))))).
14890 (* constant 2989 *)
14891 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))))).
14892 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))))).
14895 (* constant 2990 *)
14896 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)))))).
14897 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))))))).
14900 (* constant 2991 *)
14901 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)))).
14902 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))))).
14905 (* constant 2992 *)
14906 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)).
14907 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))).
14910 (* constant 2993 *)
14911 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).
14912 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)).
14915 (* constant 2994 *)
14916 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)))).
14917 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)))).
14920 (* constant 2995 *)
14921 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)))).
14922 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))))).
14925 (* constant 2996 *)
14926 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)))).
14927 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)).
14930 (* constant 2997 *)
14931 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))).
14932 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))).
14935 (* constant 2998 *)
14936 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))))).
14937 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))))).
14940 (* constant 2999 *)
14941 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)))))).
14942 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)))))).
14945 (* constant 3000 *)
14946 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)))))))).
14947 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)))))))).
14950 (* constant 3001 *)
14951 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)))))))))).
14952 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)))))))))).
14955 (* constant 3002 *)
14956 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)))))))))).
14957 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))))))))))).
14960 (* constant 3003 *)
14961 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)))))))))))).
14962 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)))))))))))).
14965 (* constant 3004 *)
14966 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)))))))))))).
14967 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)))))))))))).
14970 (* constant 3005 *)
14971 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)))))))))))).
14972 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)))))))))))).
14975 (* constant 3006 *)
14976 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)))))))))))).
14977 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))))))))))))).
14980 (* constant 3007 *)
14981 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)))))))))))).
14982 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)))))))))))).
14985 (* constant 3008 *)
14986 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)))))))))))).
14987 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)))))))))))).
14990 (* constant 3009 *)
14991 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)))))))))))).
14992 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)))))))))))).
14995 (* constant 3010 *)
14996 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)))))))))))).
14997 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)))))))))))).
15000 (* constant 3011 *)
15001 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))))))))))).
15002 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))))))))))).
15005 (* constant 3012 *)
15006 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))))))))))).
15007 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))))))))))).
15010 (* constant 3013 *)
15011 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))))))))))).
15012 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))))))))))).
15015 (* constant 3014 *)
15016 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))))))))))).
15017 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))))))))))).
15020 (* constant 3015 *)
15021 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))))))))))).
15022 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))))))))))).
15025 (* constant 3016 *)
15026 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))))))))))).
15027 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))))))))))).
15030 (* constant 3017 *)
15031 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)))))))))))).
15032 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))))))))))).
15035 (* constant 3018 *)
15036 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)))))))))))).
15037 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)))))))))))).
15040 (* constant 3019 *)
15041 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)))))))))).
15042 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))))))))))).
15045 (* constant 3020 *)
15046 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)))))))))).
15047 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))))))))))).
15050 (* constant 3021 *)
15051 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)))))))))).
15052 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))))))))))).
15055 (* constant 3022 *)
15056 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)))))))))).
15057 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))))))))))))))).
15060 (* constant 3023 *)
15061 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))))).
15062 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)))))))))).
15065 (* constant 3024 *)
15066 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)).
15067 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))))).
15070 (* constant 3025 *)
15071 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))).
15072 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))).
15075 (* constant 3026 *)
15076 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))))))).
15077 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))))).
15080 (* constant 3027 *)
15081 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)))))).
15082 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)))))).
15085 (* constant 3028 *)
15086 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)))))))).
15087 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)))))))).
15090 (* constant 3029 *)
15091 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)))))))))).
15092 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)))))))))).
15095 (* constant 3030 *)
15096 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)))))))))).
15097 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))))))))))).
15100 (* constant 3031 *)
15101 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)))))))))))).
15102 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)))))))))))).
15105 (* constant 3032 *)
15106 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)))))))))))).
15107 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)))))))))))).
15110 (* constant 3033 *)
15111 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)))))))))))).
15112 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))))))))))).
15115 (* constant 3034 *)
15116 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)))))))))))).
15117 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)))))))))))).
15120 (* constant 3035 *)
15121 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)))))))))))).
15122 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))))))))))))).
15125 (* constant 3036 *)
15126 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)))))))))))).
15127 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)))))))))))).
15130 (* constant 3037 *)
15131 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)))))))))))).
15132 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)))))))))))).
15135 (* constant 3038 *)
15136 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)))))))))))).
15137 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))))))))))).
15140 (* constant 3039 *)
15141 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)))))))))))).
15142 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)))))))))))).
15145 (* constant 3040 *)
15146 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)))))))))))).
15147 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)))))))))))).
15150 (* constant 3041 *)
15151 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))))))))))).
15152 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))))))))))).
15155 (* constant 3042 *)
15156 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))))))))))).
15157 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))))))))))).
15160 (* constant 3043 *)
15161 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))))))))))).
15162 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))))))))))).
15165 (* constant 3044 *)
15166 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))))))))))).
15167 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)))))))))))).
15170 (* constant 3045 *)
15171 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)))))))))))).
15172 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))))))))))).
15175 (* constant 3046 *)
15176 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))))))))))).
15177 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))))))))))).
15180 (* constant 3047 *)
15181 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)))))))))).
15182 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))))))))))).
15185 (* constant 3048 *)
15186 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)))))))))))).
15187 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))))))))))).
15190 (* constant 3049 *)
15191 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)))))))))))).
15192 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))))))))))).
15195 (* constant 3050 *)
15196 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)))))))))).
15197 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))))))))))).
15200 (* constant 3051 *)
15201 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)))))))))).
15202 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)))))))))).
15205 (* constant 3052 *)
15206 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))))).
15207 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)))))))))).
15210 (* constant 3053 *)
15211 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)).
15212 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))))).
15215 (* constant 3054 *)
15216 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).
15217 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)).
15220 (* constant 3055 *)
15221 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)).
15222 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)).
15225 (* constant 3056 *)
15226 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)).
15227 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)).
15230 (* constant 3057 *)
15231 Definition l_e_st_eq_landau_n_rt_rp_irratrp : (forall (ksi:l_e_st_eq_landau_n_rt_cut), Prop).
15232 exact (fun (ksi:l_e_st_eq_landau_n_rt_cut) => l_not (l_e_st_eq_landau_n_rt_rp_ratrp ksi)).
15235 (* constant 3058 *)
15236 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)).
15237 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)).
15240 (* constant 3059 *)
15241 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)).
15242 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)).
15245 (* constant 3060 *)
15246 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))).
15247 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))))).
15250 (* constant 3061 *)
15251 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)))).
15252 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)))).
15255 (* constant 3062 *)
15256 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)))).
15257 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)))).
15260 (* constant 3063 *)
15261 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))))).
15262 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))))).
15265 (* constant 3064 *)
15266 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)))).
15267 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)))).
15270 (* constant 3065 *)
15271 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)))).
15272 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))).
15275 (* constant 3066 *)
15276 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)).
15277 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))).
15280 (* constant 3067 *)
15281 Definition l_e_st_eq_landau_n_5162_prop2 : (forall (v:l_e_st_eq_landau_n_nat), Prop).
15282 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)).
15285 (* constant 3068 *)
15286 Definition l_e_st_eq_landau_n_5162_prop3 : Prop.
15287 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)).
15290 (* constant 3069 *)
15291 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).
15292 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)).
15295 (* constant 3070 *)
15296 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)).
15297 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)).
15300 (* constant 3071 *)
15301 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)).
15302 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)).
15305 (* constant 3072 *)
15306 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)).
15307 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)).
15310 (* constant 3073 *)
15311 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)))))).
15312 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))))).
15315 (* constant 3074 *)
15316 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)))).
15317 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))))).
15320 (* constant 3075 *)
15321 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)))).
15322 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)))))).
15325 (* constant 3076 *)
15326 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)))))).
15327 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)))))).
15330 (* constant 3077 *)
15331 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))).
15332 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)))).
15335 (* constant 3078 *)
15336 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))))).
15337 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)))).
15340 (* constant 3079 *)
15341 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))))))).
15342 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)))))).
15345 (* constant 3080 *)
15346 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)))))).
15347 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)))))).
15350 (* constant 3081 *)
15351 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)))))))).
15352 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))))))).
15355 (* constant 3082 *)
15356 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))))))))).
15357 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)))))))).
15360 (* constant 3083 *)
15361 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)))))))))).
15362 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))))))))).
15365 (* constant 3084 *)
15366 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))))))))).
15367 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)))))))).
15370 (* constant 3085 *)
15371 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)))))))))).
15372 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))))))))).
15375 (* constant 3086 *)
15376 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)))))))))).
15377 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)))))))))).
15380 (* constant 3087 *)
15381 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))))))))))).
15382 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)))))))))).
15385 (* constant 3088 *)
15386 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))))))))).
15387 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))))))))).
15390 (* constant 3089 *)
15391 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)))))))))).
15392 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)))))))))).
15395 (* constant 3090 *)
15396 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)))))))).
15397 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)))))))).
15400 (* constant 3091 *)
15401 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)))))))))).
15402 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))))))))).
15405 (* constant 3092 *)
15406 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))))))))).
15407 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)))))))).
15410 (* constant 3093 *)
15411 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))))))))).
15412 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)))))))).
15415 (* constant 3094 *)
15416 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))))))).
15417 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)))))))).
15420 (* constant 3095 *)
15421 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))))))).
15422 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)))))))).
15425 (* constant 3096 *)
15426 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))))))).
15427 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)))))))).
15430 (* constant 3097 *)
15431 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))))))).
15432 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)))))))).
15435 (* constant 3098 *)
15436 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))))).
15437 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))))))).
15440 (* constant 3099 *)
15441 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))).
15442 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))))).
15445 (* constant 3100 *)
15446 Definition l_e_st_eq_landau_n_5162_t37 : (forall (p:l_e_st_eq_landau_n_5162_prop3), l_con).
15447 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))).
15450 (* constant 3101 *)
15451 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))))).
15452 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)))).
15455 (* constant 3102 *)
15456 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)))).
15457 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))))).
15460 (* constant 3103 *)
15461 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))))).
15462 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))))).
15465 (* constant 3104 *)
15466 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))))).
15467 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))))).
15470 (* constant 3105 *)
15471 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))))).
15472 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))))).
15475 (* constant 3106 *)
15476 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)))).
15477 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))))).
15480 (* constant 3107 *)
15481 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)))).
15482 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))))).
15485 (* constant 3108 *)
15486 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)).
15487 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)))).
15490 (* constant 3109 *)
15491 Definition l_e_st_eq_landau_n_rt_rp_5162_ksi : l_e_st_eq_landau_n_rt_cut.
15492 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)))).
15495 (* constant 3110 *)
15496 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)).
15497 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)))).
15500 (* constant 3111 *)
15501 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).
15502 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).
15505 (* constant 3112 *)
15506 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))).
15507 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).
15510 (* constant 3113 *)
15511 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))).
15512 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)).
15515 (* constant 3114 *)
15516 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).
15517 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)).
15520 (* constant 3115 *)
15521 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).
15522 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)).
15525 (* constant 3116 *)
15526 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).
15527 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)).
15530 (* constant 3117 *)
15531 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).
15532 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)).
15535 (* constant 3118 *)
15536 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)))).
15537 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)))).
15540 (* constant 3119 *)
15541 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)))).
15542 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))).
15545 (* constant 3120 *)
15546 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)))).
15547 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)))).
15550 (* constant 3121 *)
15551 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)))).
15552 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)))).
15555 (* constant 3122 *)
15556 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)))))).
15557 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)))).
15560 (* constant 3123 *)
15561 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)))))).
15562 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)))).
15565 (* constant 3124 *)
15566 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))))))).
15567 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))))).
15570 (* constant 3125 *)
15571 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)))).
15572 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))))).
15575 (* constant 3126 *)
15576 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)))).
15577 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))))).
15580 (* constant 3127 *)
15581 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))))).
15582 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))))).
15585 (* constant 3128 *)
15586 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))))).
15587 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))))).
15590 (* constant 3129 *)
15591 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))))))).
15592 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)))))).
15595 (* constant 3130 *)
15596 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)))))).
15597 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))))).
15600 (* constant 3131 *)
15601 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)))).
15602 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))))).
15605 (* constant 3132 *)
15606 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)))))).
15607 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))))).
15610 (* constant 3133 *)
15611 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)))))).
15612 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))))).
15615 (* constant 3134 *)
15616 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)))))).
15617 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))))).
15620 (* constant 3135 *)
15621 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))))).
15622 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))))).
15625 (* constant 3136 *)
15626 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))))))).
15627 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))))).
15630 (* constant 3137 *)
15631 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))))))).
15632 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)))))).
15635 (* constant 3138 *)
15636 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)))))).
15637 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))))).
15640 (* constant 3139 *)
15641 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)))).
15642 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))))).
15645 (* constant 3140 *)
15646 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)))))).
15647 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))))).
15650 (* constant 3141 *)
15651 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)))))).
15652 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))))).
15655 (* constant 3142 *)
15656 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)))))).
15657 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))))).
15660 (* constant 3143 *)
15661 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))))).
15662 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))))).
15665 (* constant 3144 *)
15666 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))))))).
15667 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))))).
15670 (* constant 3145 *)
15671 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)))))).
15672 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)))))).
15675 (* constant 3146 *)
15676 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)))))))).
15677 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)))))).
15680 (* constant 3147 *)
15681 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)))))))).
15682 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))))))).
15685 (* constant 3148 *)
15686 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))))))).
15687 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)))))).
15690 (* constant 3149 *)
15691 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))))).
15692 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)))))).
15695 (* constant 3150 *)
15696 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))))))).
15697 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)))))).
15700 (* constant 3151 *)
15701 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))))))).
15702 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)))))).
15705 (* constant 3152 *)
15706 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))))))).
15707 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)))))).
15710 (* constant 3153 *)
15711 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)))))).
15712 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)))))).
15715 (* constant 3154 *)
15716 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))))).
15717 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)))).
15720 (* constant 3155 *)
15721 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)))))).
15722 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)))))).
15725 (* constant 3156 *)
15726 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))))).
15727 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)))).
15730 (* constant 3157 *)
15731 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)))))).
15732 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)))))).
15735 (* constant 3158 *)
15736 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)))).
15737 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)))).
15740 (* constant 3159 *)
15741 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)).
15742 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))).
15745 (* constant 3160 *)
15746 Definition l_e_st_eq_landau_n_rt_rp_dif : Type.
15747 exact (l_e_st_eq_landau_n_pair1type l_e_st_eq_landau_n_rt_cut).
15750 (* constant 3161 *)
15751 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)).
15752 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)).
15755 (* constant 3162 *)
15756 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).
15757 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).
15760 (* constant 3163 *)
15761 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).
15762 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).
15765 (* constant 3164 *)
15766 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)).
15767 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)).
15770 (* constant 3165 *)
15771 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)))).
15772 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)).
15775 (* constant 3166 *)
15776 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)).
15777 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)).
15780 (* constant 3167 *)
15781 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)))).
15782 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)).
15785 (* constant 3168 *)
15786 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).
15787 exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_stm a).
15790 (* constant 3169 *)
15791 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).
15792 exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_std a).
15795 (* constant 3170 *)
15796 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).
15797 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).
15800 (* constant 3171 *)
15801 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))).
15802 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).
15805 (* constant 3172 *)
15806 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))))))).
15807 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))))).
15810 (* constant 3173 *)
15811 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))))).
15812 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))))).
15815 (* constant 3174 *)
15816 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))))).
15817 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)))).
15820 (* constant 3175 *)
15821 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))))).
15822 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)))).
15825 (* constant 3176 *)
15826 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)))))).
15827 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)))).
15830 (* constant 3177 *)
15831 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)))).
15832 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)))).
15835 (* constant 3178 *)
15836 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))))).
15837 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)))).
15840 (* constant 3179 *)
15841 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))))).
15842 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)))).
15845 (* constant 3180 *)
15846 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))))))).
15847 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))))))).
15850 (* constant 3181 *)
15851 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)).
15852 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)).
15855 (* constant 3182 *)
15856 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)).
15857 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)).
15860 (* constant 3183 *)
15861 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)).
15862 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)))).
15865 (* constant 3184 *)
15866 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)))))).
15867 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)))))).
15870 (* constant 3185 *)
15871 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))))).
15872 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))))).
15875 (* constant 3186 *)
15876 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)))).
15877 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))))).
15880 (* constant 3187 *)
15881 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)))))).
15882 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)))))).
15885 (* constant 3188 *)
15886 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).
15887 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))).
15890 (* constant 3189 *)
15891 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).
15892 exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_satzd163 a).
15895 (* constant 3190 *)
15896 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))).
15897 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))).
15900 (* constant 3191 *)
15901 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))).
15902 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))).
15905 (* constant 3192 *)
15906 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))))))).
15907 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))))))).
15910 (* constant 3193 *)
15911 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))))).
15912 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))))).
15915 (* constant 3194 *)
15916 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))))).
15917 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))))).
15920 (* constant 3195 *)
15921 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))).
15922 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))).
15925 (* constant 3196 *)
15926 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))).
15927 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))).
15930 (* constant 3197 *)
15931 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))).
15932 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))).
15935 (* constant 3198 *)
15936 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))).
15937 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))).
15940 (* constant 3199 *)
15941 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)))))))).
15942 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))))).
15945 (* constant 3200 *)
15946 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)))))))).
15947 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))))))).
15950 (* constant 3201 *)
15951 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))))).
15952 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)))))).
15955 (* constant 3202 *)
15956 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))))).
15957 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))))).
15960 (* constant 3203 *)
15961 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))))).
15962 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))))).
15965 (* constant 3204 *)
15966 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))))).
15967 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)))))).
15970 (* constant 3205 *)
15971 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))))))).
15972 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)))))))).
15975 (* constant 3206 *)
15976 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))))))))).
15977 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)))))))))).
15980 (* constant 3207 *)
15981 Definition l_e_st_eq_landau_n_rt_rp_posd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop).
15982 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)).
15985 (* constant 3208 *)
15986 Definition l_e_st_eq_landau_n_rt_rp_zero : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop).
15987 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)).
15990 (* constant 3209 *)
15991 Definition l_e_st_eq_landau_n_rt_rp_negd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop).
15992 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)).
15995 (* constant 3210 *)
15996 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)))).
15997 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))).
16000 (* constant 3211 *)
16001 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)))).
16002 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)))).
16005 (* constant 3212 *)
16006 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)))).
16007 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))).
16010 (* constant 3213 *)
16011 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)).
16012 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)).
16015 (* constant 3214 *)
16016 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)).
16017 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)).
16020 (* constant 3215 *)
16021 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)).
16022 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)).
16025 (* constant 3216 *)
16026 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))))).
16027 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))))).
16030 (* constant 3217 *)
16031 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))).
16032 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)).
16035 (* constant 3218 *)
16036 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))).
16037 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)).
16040 (* constant 3219 *)
16041 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))).
16042 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)).
16045 (* constant 3220 *)
16046 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))).
16047 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)).
16050 (* constant 3221 *)
16051 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))).
16052 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)).
16055 (* constant 3222 *)
16056 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))).
16057 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)).
16060 (* constant 3223 *)
16061 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)))))).
16062 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))))).
16065 (* constant 3224 *)
16066 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)))).
16067 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))))).
16070 (* constant 3225 *)
16071 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)))))).
16072 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)))))).
16075 (* constant 3226 *)
16076 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)))).
16077 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))))).
16080 (* constant 3227 *)
16081 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)))))).
16082 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))))).
16085 (* constant 3228 *)
16086 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)))).
16087 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))))).
16090 (* constant 3229 *)
16091 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)))).
16092 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)))))).
16095 (* constant 3230 *)
16096 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).
16097 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).
16100 (* constant 3231 *)
16101 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).
16102 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)).
16105 (* constant 3232 *)
16106 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)))).
16107 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)))).
16110 (* constant 3233 *)
16111 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)))).
16112 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)))).
16115 (* constant 3234 *)
16116 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)))).
16117 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)))).
16120 (* constant 3235 *)
16121 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))).
16122 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)))).
16125 (* constant 3236 *)
16126 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)))).
16127 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)))).
16130 (* constant 3237 *)
16131 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))).
16132 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))))).
16135 (* constant 3238 *)
16136 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)).
16137 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))).
16140 (* constant 3239 *)
16141 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)).
16142 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))).
16145 (* constant 3240 *)
16146 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)).
16147 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)).
16150 (* constant 3241 *)
16151 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)))).
16152 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)))).
16155 (* constant 3242 *)
16156 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)))).
16157 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))).
16160 (* constant 3243 *)
16161 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)).
16162 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))).
16165 (* constant 3244 *)
16166 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)).
16167 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))).
16170 (* constant 3245 *)
16171 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)))).
16172 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))).
16175 (* constant 3246 *)
16176 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)))).
16177 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))).
16180 (* constant 3247 *)
16181 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)).
16182 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))).
16185 (* constant 3248 *)
16186 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))))))).
16187 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)))))).
16190 (* constant 3249 *)
16191 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)))))).
16192 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)))))).
16195 (* constant 3250 *)
16196 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))))))).
16197 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))))).
16200 (* constant 3251 *)
16201 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))))).
16202 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)))))).
16205 (* constant 3252 *)
16206 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))))))).
16207 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)))))).
16210 (* constant 3253 *)
16211 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)))))).
16212 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)))))).
16215 (* constant 3254 *)
16216 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))))))).
16217 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))))).
16220 (* constant 3255 *)
16221 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))))).
16222 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)))))).
16225 (* constant 3256 *)
16226 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)))).
16227 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)).
16230 (* constant 3257 *)
16231 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))).
16232 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)).
16235 (* constant 3258 *)
16236 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).
16237 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)).
16240 (* constant 3259 *)
16241 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)))).
16242 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)).
16245 (* constant 3260 *)
16246 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))).
16247 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)).
16250 (* constant 3261 *)
16251 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).
16252 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)).
16255 (* constant 3262 *)
16256 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))))).
16257 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))))).
16260 (* constant 3263 *)
16261 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)))).
16262 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)))).
16265 (* constant 3264 *)
16266 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)))).
16267 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))).
16270 (* constant 3265 *)
16271 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)).
16272 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))).
16275 (* constant 3266 *)
16276 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).
16277 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).
16280 (* constant 3267 *)
16281 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)))).
16282 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))).
16285 (* constant 3268 *)
16286 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)).
16287 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))).
16290 (* constant 3269 *)
16291 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)))).
16292 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))))).
16295 (* constant 3270 *)
16296 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)))).
16297 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))))).
16300 (* constant 3271 *)
16301 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)))))).
16302 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)))))).
16305 (* constant 3272 *)
16306 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))))).
16307 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))))))).
16310 (* constant 3273 *)
16311 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))))).
16312 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)))))))).
16315 (* constant 3274 *)
16316 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)))).
16317 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)))).
16320 (* constant 3275 *)
16321 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))).
16322 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)).
16325 (* constant 3276 *)
16326 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)))).
16327 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))).
16330 (* constant 3277 *)
16331 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))).
16332 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))).
16335 (* constant 3278 *)
16336 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))))).
16337 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))))))).
16340 (* constant 3279 *)
16341 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))))))).
16342 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)))))).
16345 (* constant 3280 *)
16346 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))))).
16347 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))))))).
16350 (* constant 3281 *)
16351 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))).
16352 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))).
16355 (* constant 3282 *)
16356 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))).
16357 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)).
16360 (* constant 3283 *)
16361 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)).
16362 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)))).
16365 (* constant 3284 *)
16366 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)))))).
16367 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))))).
16370 (* constant 3285 *)
16371 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))))).
16372 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)))).
16375 (* constant 3286 *)
16376 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)))).
16377 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)))).
16380 (* constant 3287 *)
16381 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)))))).
16382 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))))).
16385 (* constant 3288 *)
16386 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)).
16387 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)))).
16390 (* constant 3289 *)
16391 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))).
16392 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))).
16395 (* constant 3290 *)
16396 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))).
16397 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))).
16400 (* constant 3291 *)
16401 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)))))).
16402 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))))))).
16405 (* constant 3292 *)
16406 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))))).
16407 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)))))).
16410 (* constant 3293 *)
16411 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)))).
16412 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)))))).
16415 (* constant 3294 *)
16416 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)))))).
16417 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))))))).
16420 (* constant 3295 *)
16421 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))).
16422 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)))).
16425 (* constant 3296 *)
16426 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))).
16427 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)))).
16430 (* constant 3297 *)
16431 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))).
16432 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)))).
16435 (* constant 3298 *)
16436 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)))).
16437 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)))).
16440 (* constant 3299 *)
16441 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)))).
16442 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)))).
16445 (* constant 3300 *)
16446 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)))))))))).
16447 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)))))))))).
16450 (* constant 3301 *)
16451 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)))))))))).
16452 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)))))))).
16455 (* constant 3302 *)
16456 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))))))).
16457 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)))))))).
16460 (* constant 3303 *)
16461 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))))))).
16462 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))))))))).
16465 (* constant 3304 *)
16466 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))))).
16467 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))))).
16470 (* constant 3305 *)
16471 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))))).
16472 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))))).
16475 (* constant 3306 *)
16476 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))))).
16477 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))))).
16480 (* constant 3307 *)
16481 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))))).
16482 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))))).
16485 (* constant 3308 *)
16486 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)).
16487 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))).
16490 (* constant 3309 *)
16491 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)).
16492 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))).
16495 (* constant 3310 *)
16496 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))).
16497 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)))).
16500 (* constant 3311 *)
16501 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))).
16502 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)))).
16505 (* constant 3312 *)
16506 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))))).
16507 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)))))).
16510 (* constant 3313 *)
16511 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))))).
16512 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)))))).
16515 (* constant 3314 *)
16516 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))))).
16517 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))))))).
16520 (* constant 3315 *)
16521 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))))).
16522 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))))))).
16525 (* constant 3316 *)
16526 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))))))).
16527 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)))))))).
16530 (* constant 3317 *)
16531 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))))))).
16532 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)))))))).
16535 (* constant 3318 *)
16536 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))).
16537 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))).
16540 (* constant 3319 *)
16541 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))).
16542 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))).
16545 (* constant 3320 *)
16546 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))).
16547 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))).
16550 (* constant 3321 *)
16551 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))).
16552 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))).
16555 (* constant 3322 *)
16556 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)))).
16557 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)))).
16560 (* constant 3323 *)
16561 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)))).
16562 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))).
16565 (* constant 3324 *)
16566 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))).
16567 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))).
16570 (* constant 3325 *)
16571 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))).
16572 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)))).
16575 (* constant 3326 *)
16576 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)))).
16577 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)))).
16580 (* constant 3327 *)
16581 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)))).
16582 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)))).
16585 (* constant 3328 *)
16586 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))).
16587 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)))).
16590 (* constant 3329 *)
16591 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))).
16592 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)))).
16595 (* constant 3330 *)
16596 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)))).
16597 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))))).
16600 (* constant 3331 *)
16601 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)))).
16602 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))))).
16605 (* constant 3332 *)
16606 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)))).
16607 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))))).
16610 (* constant 3333 *)
16611 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)))).
16612 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))))).
16615 (* constant 3334 *)
16616 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)))).
16617 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)))))).
16620 (* constant 3335 *)
16621 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)))).
16622 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)))))).
16625 (* constant 3336 *)
16626 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)))).
16627 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)))))).
16630 (* constant 3337 *)
16631 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))).
16632 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)))).
16635 (* constant 3338 *)
16636 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)))))))).
16637 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))))).
16640 (* constant 3339 *)
16641 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)))))))).
16642 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)))))).
16645 (* constant 3340 *)
16646 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))))).
16647 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)))))).
16650 (* constant 3341 *)
16651 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))))).
16652 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))))).
16655 (* constant 3342 *)
16656 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))))).
16657 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))))))).
16660 (* constant 3343 *)
16661 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))))).
16662 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)))))).
16665 (* constant 3344 *)
16666 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))))).
16667 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)))))).
16670 (* constant 3345 *)
16671 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))))).
16672 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))))))).
16675 (* constant 3346 *)
16676 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))))).
16677 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))))))).
16680 (* constant 3347 *)
16681 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)))))).
16682 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))))))).
16685 (* constant 3348 *)
16686 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)))))).
16687 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)))))).
16690 (* constant 3349 *)
16691 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))))).
16692 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)))))).
16695 (* constant 3350 *)
16696 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))))).
16697 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))))).
16700 (* constant 3351 *)
16701 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))))).
16702 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))))))).
16705 (* constant 3352 *)
16706 Definition l_e_st_eq_landau_n_rt_rp_ratd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop).
16707 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)))).
16710 (* constant 3353 *)
16711 Definition l_e_st_eq_landau_n_rt_rp_irratd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop).
16712 exact (fun (a:l_e_st_eq_landau_n_rt_rp_dif) => l_not (l_e_st_eq_landau_n_rt_rp_ratd a)).
16715 (* constant 3354 *)
16716 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)))))).
16717 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)))))).
16720 (* constant 3355 *)
16721 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))))))).
16722 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)))))).
16725 (* constant 3356 *)
16726 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))))))).
16727 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)))))).
16730 (* constant 3357 *)
16731 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)))).
16732 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))))).
16735 (* constant 3358 *)
16736 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)))).
16737 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))))).
16740 (* constant 3359 *)
16741 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)).
16742 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))).
16745 (* constant 3360 *)
16746 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))))))).
16747 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)))))).
16750 (* constant 3361 *)
16751 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))))))).
16752 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))))))).
16755 (* constant 3362 *)
16756 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)))))).
16757 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))))))).
16760 (* constant 3363 *)
16761 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))))))))).
16762 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))))))).
16765 (* constant 3364 *)
16766 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))))))))).
16767 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))))))).
16770 (* constant 3365 *)
16771 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))))))))).
16772 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))))))).
16775 (* constant 3366 *)
16776 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)))))).
16777 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))))))).
16780 (* constant 3367 *)
16781 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)))).
16782 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))))))).
16785 (* constant 3368 *)
16786 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))))).
16787 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)))).
16790 (* constant 3369 *)
16791 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).
16792 exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_pdofrp r).
16795 (* constant 3370 *)
16796 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).
16797 exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_ndofrp r).
16800 (* constant 3371 *)
16801 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)).
16802 exact (fun (r:l_e_st_eq_landau_n_rt_cut) => l_e_st_eq_landau_n_rt_rp_posdirp r).
16805 (* constant 3372 *)
16806 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))).
16807 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)).
16810 (* constant 3373 *)
16811 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))).
16812 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)).
16815 (* constant 3374 *)
16816 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)))).
16817 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)))))).
16820 (* constant 3375 *)
16821 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)).
16822 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)))).
16825 (* constant 3376 *)
16826 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)))).
16827 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)))).
16830 (* constant 3377 *)
16831 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)))).
16832 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)))).
16835 (* constant 3378 *)
16836 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)))).
16837 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)))).
16840 (* constant 3379 *)
16841 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))).
16842 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))).
16845 (* constant 3380 *)
16846 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))).
16847 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)).
16850 (* constant 3381 *)
16851 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))).
16852 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))).
16855 (* constant 3382 *)
16856 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))).
16857 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)).
16860 (* constant 3383 *)
16861 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))).
16862 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))))).
16865 (* constant 3384 *)
16866 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))).
16867 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)).
16870 (* constant 3385 *)
16871 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))).
16872 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))))).
16875 (* constant 3386 *)
16876 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))).
16877 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)).
16880 (* constant 3387 *)
16881 Definition l_e_st_eq_landau_n_rt_rp_natd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop).
16882 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))).
16885 (* constant 3388 *)
16886 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)).
16887 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)).
16890 (* constant 3389 *)
16891 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)))).
16892 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)).
16895 (* constant 3390 *)
16896 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)))).
16897 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))))).
16900 (* constant 3391 *)
16901 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)))))).
16902 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))))).
16905 (* constant 3392 *)
16906 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)))))).
16907 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)))))).
16910 (* constant 3393 *)
16911 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)))).
16912 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))))).
16915 (* constant 3394 *)
16916 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).
16917 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)).
16920 (* constant 3395 *)
16921 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)).
16922 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)).
16925 (* constant 3396 *)
16926 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)))).
16927 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))).
16930 (* constant 3397 *)
16931 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))).
16932 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)))).
16935 (* constant 3398 *)
16936 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))).
16937 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))).
16940 (* constant 3399 *)
16941 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))).
16942 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))).
16945 (* constant 3400 *)
16946 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)).
16947 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)).
16950 (* constant 3401 *)
16951 Definition l_e_st_eq_landau_n_rt_rp_intd : (forall (a:l_e_st_eq_landau_n_rt_rp_dif), Prop).
16952 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))).
16955 (* constant 3402 *)
16956 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))))).
16957 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))))).
16960 (* constant 3403 *)
16961 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)))))).
16962 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))))).
16965 (* constant 3404 *)
16966 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)))).
16967 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))))).
16970 (* constant 3405 *)
16971 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))).
16972 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))))).
16975 (* constant 3406 *)
16976 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))).
16977 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)).
16980 (* constant 3407 *)
16981 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)).
16982 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))).
16985 (* constant 3408 *)
16986 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)))).
16987 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)))).
16990 (* constant 3409 *)
16991 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))).
16992 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)))).
16995 (* constant 3410 *)
16996 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)).
16997 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)).
17000 (* constant 3411 *)
17001 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))).
17002 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)).
17005 (* constant 3412 *)
17006 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)))).
17007 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)))))).
17010 (* constant 3413 *)
17011 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)))).
17012 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)))).
17015 (* constant 3414 *)
17016 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))).
17017 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))).
17020 (* constant 3415 *)
17021 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))).
17022 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))).
17025 (* constant 3416 *)
17026 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))).
17027 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)))).
17030 (* constant 3417 *)
17031 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)))).
17032 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))).
17035 (* constant 3418 *)
17036 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))).
17037 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))).
17040 (* constant 3419 *)
17041 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)))).
17042 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))).
17045 (* constant 3420 *)
17046 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))))).
17047 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)))).
17050 (* constant 3421 *)
17051 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))))).
17052 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))))).
17055 (* constant 3422 *)
17056 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)).
17057 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))).
17060 (* constant 3423 *)
17061 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)).
17062 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)))).
17065 (* constant 3424 *)
17066 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)))))).
17067 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)))))).
17070 (* constant 3425 *)
17071 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))))).
17072 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))))).
17075 (* constant 3426 *)
17076 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)))))).
17077 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))))).
17080 (* constant 3427 *)
17081 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)))))).
17082 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))))).
17085 (* constant 3428 *)
17086 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)))))).
17087 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))))).
17090 (* constant 3429 *)
17091 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))))).
17092 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)))).
17095 (* constant 3430 *)
17096 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))))).
17097 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)))).
17100 (* constant 3431 *)
17101 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)))))).
17102 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)))).
17105 (* constant 3432 *)
17106 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)))).
17107 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)))).
17110 (* constant 3433 *)
17111 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))).
17112 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)))).
17115 (* constant 3434 *)
17116 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))).
17117 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)).
17120 (* constant 3435 *)
17121 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))))))).
17122 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)))))).
17125 (* constant 3436 *)
17126 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))))).
17127 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))))).
17130 (* constant 3437 *)
17131 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))))).
17132 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))))).
17135 (* constant 3438 *)
17136 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))))))).
17137 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))))))).
17140 (* constant 3439 *)
17141 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)))))).
17142 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))))).
17145 (* constant 3440 *)
17146 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))).
17147 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)))).
17150 (* constant 3441 *)
17151 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))).
17152 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)))).
17155 (* constant 3442 *)
17156 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))))).
17157 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))))).
17160 (* constant 3443 *)
17161 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))))).
17162 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))))).
17165 (* constant 3444 *)
17166 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).
17167 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)).
17170 (* constant 3445 *)
17171 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))).
17172 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))).
17175 (* constant 3446 *)
17176 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)))).
17177 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))).
17180 (* constant 3447 *)
17181 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))))).
17182 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))))).
17185 (* constant 3448 *)
17186 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)))).
17187 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)))).
17190 (* constant 3449 *)
17191 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))).
17192 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))).
17195 (* constant 3450 *)
17196 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))).
17197 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))).
17200 (* constant 3451 *)
17201 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))).
17202 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))).
17205 (* constant 3452 *)
17206 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)).
17207 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))).
17210 (* constant 3453 *)
17211 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)).
17212 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))))).
17215 (* constant 3454 *)
17216 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)).
17217 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))).
17220 (* constant 3455 *)
17221 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)).
17222 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)).
17225 (* constant 3456 *)
17226 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).
17227 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)).
17230 (* constant 3457 *)
17231 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).
17232 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)).
17235 (* constant 3458 *)
17236 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))).
17237 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)).
17240 (* constant 3459 *)
17241 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))).
17242 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)))).
17245 (* constant 3460 *)
17246 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)))).
17247 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)))).
17250 (* constant 3461 *)
17251 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)))).
17252 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)))).
17255 (* constant 3462 *)
17256 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))).
17257 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)))).
17260 (* constant 3463 *)
17261 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))).
17262 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))))).
17265 (* constant 3464 *)
17266 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))).
17267 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))))).
17270 (* constant 3465 *)
17271 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))).
17272 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)))).
17275 (* constant 3466 *)
17276 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)).
17277 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)).
17280 (* constant 3467 *)
17281 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))).
17282 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)).
17285 (* constant 3468 *)
17286 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))).
17287 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)).
17290 (* constant 3469 *)
17291 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)))).
17292 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))).
17295 (* constant 3470 *)
17296 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))).
17297 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)).
17300 (* constant 3471 *)
17301 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)).
17302 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)).
17305 (* constant 3472 *)
17306 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)))).
17307 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)))).
17310 (* constant 3473 *)
17311 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)))).
17312 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))).
17315 (* constant 3474 *)
17316 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)).
17317 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))).
17320 (* constant 3475 *)
17321 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)))))).
17322 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))))).
17325 (* constant 3476 *)
17326 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)))))).
17327 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))))).
17330 (* constant 3477 *)
17331 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))))).
17332 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)))).
17335 (* constant 3478 *)
17336 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))))).
17337 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)))).
17340 (* constant 3479 *)
17341 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)))))).
17342 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)))).
17345 (* constant 3480 *)
17346 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)))).
17347 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)))).
17350 (* constant 3481 *)
17351 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))))).
17352 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)))).
17355 (* constant 3482 *)
17356 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))))).
17357 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))))).
17360 (* constant 3483 *)
17361 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))))))).
17362 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))))))).
17365 (* constant 3484 *)
17366 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))).
17367 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))).
17370 (* constant 3485 *)
17371 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)))).
17372 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))).
17375 (* constant 3486 *)
17376 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))))).
17377 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))))).
17380 (* constant 3487 *)
17381 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))).
17382 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))).
17385 (* constant 3488 *)
17386 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)))).
17387 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)))).
17390 (* constant 3489 *)
17391 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)))).
17392 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)))).
17395 (* constant 3490 *)
17396 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)))))).
17397 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))).
17400 (* constant 3491 *)
17401 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))).
17402 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)))).
17405 (* constant 3492 *)
17406 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)))))).
17407 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))).
17410 (* constant 3493 *)
17411 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))).
17412 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)))).
17415 (* constant 3494 *)
17416 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)))))).
17417 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))).
17420 (* constant 3495 *)
17421 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))).
17422 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)))).
17425 (* constant 3496 *)
17426 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)))))).
17427 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))).
17430 (* constant 3497 *)
17431 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)))).
17432 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)))).
17435 (* constant 3498 *)
17436 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)))))).
17437 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))).
17440 (* constant 3499 *)
17441 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)))).
17442 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)))).
17445 (* constant 3500 *)
17446 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)))))).
17447 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))).
17450 (* constant 3501 *)
17451 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)))).
17452 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)))).
17455 (* constant 3502 *)
17456 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))))).
17457 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)))).
17460 (* constant 3503 *)
17461 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))))).
17462 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)).
17465 (* constant 3504 *)
17466 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)))).
17467 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)))).
17470 (* constant 3505 *)
17471 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)))).
17472 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))).
17475 (* constant 3506 *)
17476 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)))).
17477 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)))).
17480 (* constant 3507 *)
17481 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))).
17482 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)))).
17485 (* constant 3508 *)
17486 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))).
17487 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)))).
17490 (* constant 3509 *)
17491 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))).
17492 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)))).
17495 (* constant 3510 *)
17496 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)))).
17497 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)).
17500 (* constant 3511 *)
17501 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))))).
17502 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)).
17505 (* constant 3512 *)
17506 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)))).
17507 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)).
17510 (* constant 3513 *)
17511 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))))).
17512 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)).
17515 (* constant 3514 *)
17516 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))))).
17517 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)))))).
17520 (* constant 3515 *)
17521 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)))).
17522 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)))).
17525 (* constant 3516 *)
17526 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)))).
17527 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)))).
17530 (* constant 3517 *)
17531 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)))))).
17532 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))))).
17535 (* constant 3518 *)
17536 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)).
17537 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)))).
17540 (* constant 3519 *)
17541 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)).
17542 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)))).
17545 (* constant 3520 *)
17546 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)))))).
17547 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)))))).
17550 (* constant 3521 *)
17551 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))))).
17552 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))).
17555 (* constant 3522 *)
17556 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)).
17557 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))).
17560 (* constant 3523 *)
17561 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)))).
17562 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))))).
17565 (* constant 3524 *)
17566 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))))).
17567 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))))).
17570 (* constant 3525 *)
17571 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)))).
17572 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))))).
17575 (* constant 3526 *)
17576 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))))).
17577 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))))).
17580 (* constant 3527 *)
17581 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)))).
17582 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))))).
17585 (* constant 3528 *)
17586 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))))).
17587 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)))).
17590 (* constant 3529 *)
17591 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))))).
17592 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))))).
17595 (* constant 3530 *)
17596 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)))).
17597 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))))).
17600 (* constant 3531 *)
17601 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))))).
17602 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))))).
17605 (* constant 3532 *)
17606 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)))).
17607 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))))).
17610 (* constant 3533 *)
17611 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))))).
17612 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))))).
17615 (* constant 3534 *)
17616 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)))).
17617 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))))).
17620 (* constant 3535 *)
17621 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)))))).
17622 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))))).
17625 (* constant 3536 *)
17626 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))))).
17627 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))))).
17630 (* constant 3537 *)
17631 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))))).
17632 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)))).
17635 (* constant 3538 *)
17636 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)))))).
17637 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))))).
17640 (* constant 3539 *)
17641 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))))).
17642 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))))).
17645 (* constant 3540 *)
17646 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)))).
17647 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))))).
17650 (* constant 3541 *)
17651 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)))).
17652 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)))))).
17655 (* constant 3542 *)
17656 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))))).
17657 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)))))).
17660 (* constant 3543 *)
17661 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))))).
17662 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))))).
17665 (* constant 3544 *)
17666 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))))).
17667 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)))).
17670 (* constant 3545 *)
17671 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))))).
17672 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))))).
17675 (* constant 3546 *)
17676 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))))))).
17677 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))))))).
17680 (* constant 3547 *)
17681 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))))))).
17682 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))))))).
17685 (* constant 3548 *)
17686 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))))))).
17687 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))))))).
17690 (* constant 3549 *)
17691 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))))))).
17692 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))))))).
17695 (* constant 3550 *)
17696 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))))))).
17697 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))))))).
17700 (* constant 3551 *)
17701 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))))))).
17702 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)))))))).
17705 (* constant 3552 *)
17706 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))))))).
17707 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))))))).
17710 (* constant 3553 *)
17711 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))))))).
17712 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))))))).
17715 (* constant 3554 *)
17716 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))))))).
17717 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)))))))).
17720 (* constant 3555 *)
17721 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))))))).
17722 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)))))))).
17725 (* constant 3556 *)
17726 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))))))))).
17727 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))))))))).
17730 (* constant 3557 *)
17731 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))))))))).
17732 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))))))))).
17735 (* constant 3558 *)
17736 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)))))))).
17737 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)))))))).
17740 (* constant 3559 *)
17741 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)))))))).
17742 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)))))))).
17745 (* constant 3560 *)
17746 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))))))).
17747 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))))))).
17750 (* constant 3561 *)
17751 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))))))).
17752 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)))))))).
17755 (* constant 3562 *)
17756 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)).
17757 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))))).
17760 (* constant 3563 *)
17761 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)))).
17762 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)))).
17765 (* constant 3564 *)
17766 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)))).
17767 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)))).
17770 (* constant 3565 *)
17771 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)))).
17772 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)))).
17775 (* constant 3566 *)
17776 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)))).
17777 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)))).
17780 (* constant 3567 *)
17781 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)))))).
17782 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))))).
17785 (* constant 3568 *)
17786 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)))))).
17787 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))))).
17790 (* constant 3569 *)
17791 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)))))).
17792 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))))).
17795 (* constant 3570 *)
17796 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)))))).
17797 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))))).
17800 (* constant 3571 *)
17801 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)))))).
17802 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))))).
17805 (* constant 3572 *)
17806 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)))))).
17807 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))))).
17810 (* constant 3573 *)
17811 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))))))).
17812 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))))).
17815 (* constant 3574 *)
17816 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)))))).
17817 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))))).
17820 (* constant 3575 *)
17821 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))))))).
17822 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))))).
17825 (* constant 3576 *)
17826 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))))))).
17827 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))))).
17830 (* constant 3577 *)
17831 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)))))).
17832 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))))).
17835 (* constant 3578 *)
17836 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)))))).
17837 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)))).
17840 (* constant 3579 *)
17841 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))))).
17842 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)))).
17845 (* constant 3580 *)
17846 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))))))).
17847 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)))).
17850 (* constant 3581 *)
17851 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)))).
17852 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)))).
17855 (* constant 3582 *)
17856 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))))).
17857 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)))).
17860 (* constant 3583 *)
17861 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))))).
17862 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))))).
17865 (* constant 3584 *)
17866 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))).
17867 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))).
17870 (* constant 3585 *)
17871 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))).
17872 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)).
17875 (* constant 3586 *)
17876 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))))))).
17877 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)))))).
17880 (* constant 3587 *)
17881 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)))))))).
17882 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))))))).
17885 (* constant 3588 *)
17886 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))))).
17887 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))))).
17890 (* constant 3589 *)
17891 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))))).
17892 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))))).
17895 (* constant 3590 *)
17896 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))))))).
17897 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))))))).
17900 (* constant 3591 *)
17901 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)))))).
17902 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)))))).
17905 (* constant 3592 *)
17906 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)))).
17907 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)))).
17910 (* constant 3593 *)
17911 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)))).
17912 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)))).
17915 (* constant 3594 *)
17916 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)))).
17917 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))).
17920 (* constant 3595 *)
17921 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)))).
17922 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))).
17925 (* constant 3596 *)
17926 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)))).
17927 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)))))).
17930 (* constant 3597 *)
17931 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)))).
17932 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)))).
17935 (* constant 3598 *)
17936 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)))).
17937 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))).
17940 (* constant 3599 *)
17941 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))).
17942 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))).
17945 (* constant 3600 *)
17946 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))).
17947 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))).
17950 (* constant 3601 *)
17951 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)))).
17952 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))).
17955 (* constant 3602 *)
17956 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))).
17957 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)))).
17960 (* constant 3603 *)
17961 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)))).
17962 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))).
17965 (* constant 3604 *)
17966 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))))))).
17967 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))))))).
17970 (* constant 3605 *)
17971 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))))))).
17972 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)))))))).
17975 (* constant 3606 *)
17976 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))))))))).
17977 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))))))).
17980 (* constant 3607 *)
17981 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)))))))).
17982 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))))).
17985 (* constant 3608 *)
17986 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))))))).
17987 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))))).
17990 (* constant 3609 *)
17991 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))))).
17992 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))))).
17995 (* constant 3610 *)
17996 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)))))).
17997 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)))))).
18000 (* constant 3611 *)
18001 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))))).
18002 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))))).
18005 (* constant 3612 *)
18006 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))))).
18007 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))))).
18010 (* constant 3613 *)
18011 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))))).
18012 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)))))).
18015 (* constant 3614 *)
18016 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)))))))).
18017 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))))))).
18020 (* constant 3615 *)
18021 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)))))))).
18022 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))))))).
18025 (* constant 3616 *)
18026 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))))))).
18027 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)))))).
18030 (* constant 3617 *)
18031 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)))))))).
18032 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))))))).
18035 (* constant 3618 *)
18036 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)))))))).
18037 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))))))).
18040 (* constant 3619 *)
18041 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))))))).
18042 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)))))).
18045 (* constant 3620 *)
18046 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)))))).
18047 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))))).
18050 (* constant 3621 *)
18051 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)))).
18052 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)))))).
18055 (* constant 3622 *)
18056 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)))).
18057 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)))).
18060 (* constant 3623 *)
18061 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)))))).
18062 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))))))).
18065 (* constant 3624 *)
18066 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)))))).
18067 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))))).
18070 (* constant 3625 *)
18071 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)))))).
18072 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)))))).
18075 (* constant 3626 *)
18076 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))))).
18077 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))))).
18080 (* constant 3627 *)
18081 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))))).
18082 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))))).
18085 (* constant 3628 *)
18086 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)))))).
18087 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)))))).
18090 (* constant 3629 *)
18091 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))))).
18092 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))))).
18095 (* constant 3630 *)
18096 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)))))).
18097 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))))).
18100 (* constant 3631 *)
18101 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))))).
18102 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)))).
18105 (* constant 3632 *)
18106 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))))).
18107 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)))).
18110 (* constant 3633 *)
18111 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)))).
18112 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))).
18115 (* constant 3634 *)
18116 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)))).
18117 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))).
18120 (* constant 3635 *)
18121 Definition l_e_st_eq_landau_n_rt_rp_1df : l_e_st_eq_landau_n_rt_rp_dif.
18122 exact (l_e_st_eq_landau_n_rt_rp_pdofrp l_e_st_eq_landau_n_rt_rp_1rp).
18125 (* constant 3636 *)
18126 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)))).
18127 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))))).
18130 (* constant 3637 *)
18131 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)))).
18132 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)))).
18135 (* constant 3638 *)
18136 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))))).
18137 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)).
18140 (* constant 3639 *)
18141 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).
18142 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))).
18145 (* constant 3640 *)
18146 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)).
18147 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)).
18150 (* constant 3641 *)
18151 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).
18152 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)).
18155 (* constant 3642 *)
18156 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)).
18157 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)).
18160 (* constant 3643 *)
18161 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)))))).
18162 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))))))).
18165 (* constant 3644 *)
18166 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)))))).
18167 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)))))).
18170 (* constant 3645 *)
18171 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))))))).
18172 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)))))).
18175 (* constant 3646 *)
18176 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))))))).
18177 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))))))).
18180 (* constant 3647 *)
18181 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)).
18182 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))).
18185 (* constant 3648 *)
18186 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)).
18187 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))).
18190 (* constant 3649 *)
18191 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)).
18192 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))).
18195 (* constant 3650 *)
18196 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)).
18197 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))).
18200 (* constant 3651 *)
18201 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)))))).
18202 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))))).
18205 (* constant 3652 *)
18206 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))))))).
18207 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))))))).
18210 (* constant 3653 *)
18211 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))))))).
18212 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))))))).
18215 (* constant 3654 *)
18216 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)))))).
18217 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)))))).
18220 (* constant 3655 *)
18221 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)))))).
18222 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))))))).
18225 (* constant 3656 *)
18226 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))))))).
18227 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))))))).
18230 (* constant 3657 *)
18231 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))))))).
18232 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))))))).
18235 (* constant 3658 *)
18236 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)))))).
18237 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))))))).
18240 (* constant 3659 *)
18241 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)))))).
18242 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))))))).
18245 (* constant 3660 *)
18246 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))))))).
18247 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))))))).
18250 (* constant 3661 *)
18251 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))))))).
18252 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)))))).
18255 (* constant 3662 *)
18256 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))))))).
18257 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))))).
18260 (* constant 3663 *)
18261 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))))))).
18262 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))))))).
18265 (* constant 3664 *)
18266 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))))))).
18267 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))))))).
18270 (* constant 3665 *)
18271 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)))))).
18272 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))))))).
18275 (* constant 3666 *)
18276 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)))))).
18277 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))))))).
18280 (* constant 3667 *)
18281 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))))))).
18282 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))))))).
18285 (* constant 3668 *)
18286 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))))))).
18287 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))))))).
18290 (* constant 3669 *)
18291 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)))))).
18292 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)))))).
18295 (* constant 3670 *)
18296 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)))))).
18297 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))))))).
18300 (* constant 3671 *)
18301 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))))))).
18302 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))))))).
18305 (* constant 3672 *)
18306 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))))))).
18307 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)))))).
18310 (* constant 3673 *)
18311 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))))))))).
18312 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)))))))).
18315 (* constant 3674 *)
18316 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))))))))).
18317 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)))))))).
18320 (* constant 3675 *)
18321 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)))))))))).
18322 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)))))))).
18325 (* constant 3676 *)
18326 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))))).
18327 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))))))).
18330 (* constant 3677 *)
18331 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))))).
18332 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))).
18335 (* constant 3678 *)
18336 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)))).
18337 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)))).
18340 (* constant 3679 *)
18341 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))))))))).
18342 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)))))))).
18345 (* constant 3680 *)
18346 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))))).
18347 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))))))).
18350 (* constant 3681 *)
18351 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))))).
18352 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))))).
18355 (* constant 3682 *)
18356 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))))).
18357 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))).
18360 (* constant 3683 *)
18361 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)))).
18362 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)))).
18365 (* constant 3684 *)
18366 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))))).
18367 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)))).
18370 (* constant 3685 *)
18371 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))))).
18372 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))))).
18375 (* constant 3686 *)
18376 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))))).
18377 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))))).
18380 (* constant 3687 *)
18381 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))))).
18382 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))).
18385 (* constant 3688 *)
18386 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)))).
18387 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)))).
18390 (* constant 3689 *)
18391 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))))).
18392 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)))).
18395 (* constant 3690 *)
18396 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))))).
18397 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))).
18400 (* constant 3691 *)
18401 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))))).
18402 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)))).
18405 (* constant 3692 *)
18406 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))))))).
18407 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)))))).
18410 (* constant 3693 *)
18411 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)))))).
18412 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)))))).
18415 (* constant 3694 *)
18416 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)))))).
18417 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)))))).
18420 (* constant 3695 *)
18421 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))))))).
18422 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)))))).
18425 (* constant 3696 *)
18426 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)))))).
18427 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)))))).
18430 (* constant 3697 *)
18431 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)))))).
18432 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)))))).
18435 (* constant 3698 *)
18436 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)))))).
18437 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)))))).
18440 (* constant 3699 *)
18441 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)))))).
18442 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)))))).
18445 (* constant 3700 *)
18446 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)))))).
18447 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)))))).
18450 (* constant 3701 *)
18451 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)))))).
18452 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)))))).
18455 (* constant 3702 *)
18456 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)))))).
18457 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)))))).
18460 (* constant 3703 *)
18461 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)))))).
18462 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)))))).
18465 (* constant 3704 *)
18466 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)))))).
18467 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)))))).
18470 (* constant 3705 *)
18471 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)))))).
18472 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)))))).
18475 (* constant 3706 *)
18476 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))).
18477 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)))).
18480 (* constant 3707 *)
18481 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))))).
18482 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)))).
18485 (* constant 3708 *)
18486 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))))).
18487 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))))).
18490 (* constant 3709 *)
18491 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)).
18492 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)).
18495 (* constant 3710 *)
18496 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)).
18497 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))).
18500 (* constant 3711 *)
18501 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)).
18502 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))).
18505 (* constant 3712 *)
18506 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))))).
18507 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))))).
18510 (* constant 3713 *)
18511 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))).
18512 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))))).
18515 (* constant 3714 *)
18516 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))))).
18517 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)))).
18520 (* constant 3715 *)
18521 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)).
18522 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))).
18525 (* constant 3716 *)
18526 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)).
18527 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))).
18530 (* constant 3717 *)
18531 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))).
18532 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))).
18535 (* constant 3718 *)
18536 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))).
18537 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)).
18540 (* constant 3719 *)
18541 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)))).
18542 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)))).
18545 (* constant 3720 *)
18546 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))))).
18547 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))))).
18550 (* constant 3721 *)
18551 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))).
18552 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)))).
18555 (* constant 3722 *)
18556 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))).
18557 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))).
18560 (* constant 3723 *)
18561 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)))))))).
18562 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))))))).
18565 (* constant 3724 *)
18566 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))))))))).
18567 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))))))))).
18570 (* constant 3725 *)
18571 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)))))))).
18572 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))))))).
18575 (* constant 3726 *)
18576 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))))))).
18577 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)))))))).
18580 (* constant 3727 *)
18581 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))))).
18582 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)))))).
18585 (* constant 3728 *)
18586 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)))))).
18587 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)))))).
18590 (* constant 3729 *)
18591 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)))).
18592 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))))).
18595 (* constant 3730 *)
18596 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)))).
18597 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)))).
18600 (* constant 3731 *)
18601 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)))).
18602 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)))).
18605 (* constant 3732 *)
18606 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)))).
18607 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))).
18610 (* constant 3733 *)
18611 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))).
18612 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))))).
18615 (* constant 3734 *)
18616 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)))).
18617 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))))).
18620 (* constant 3735 *)
18621 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))).
18622 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))))).
18625 (* constant 3736 *)
18626 Definition l_e_st_eq_landau_n_rt_rp_iv5d_i : l_e_st_eq_landau_n_rt_cut.
18627 exact l_e_st_eq_landau_n_rt_rp_1rp.
18630 (* constant 3737 *)
18631 Definition l_e_st_eq_landau_n_rt_rp_iv5d_2 : l_e_st_eq_landau_n_rt_cut.
18632 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).
18635 (* constant 3738 *)
18636 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).
18637 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).
18640 (* constant 3739 *)
18641 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)).
18642 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)).
18645 (* constant 3740 *)
18646 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)).
18647 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)).
18650 (* constant 3741 *)
18651 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)).
18652 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)).
18655 (* constant 3742 *)
18656 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))).
18657 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))).
18660 (* constant 3743 *)
18661 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)))).
18662 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)))).
18665 (* constant 3744 *)
18666 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)))).
18667 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)).
18670 (* constant 3745 *)
18671 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))).
18672 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)))).
18675 (* constant 3746 *)
18676 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))).
18677 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)))).
18680 (* constant 3747 *)
18681 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)))).
18682 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))).
18685 (* constant 3748 *)
18686 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))).
18687 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))).
18690 (* constant 3749 *)
18691 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))).
18692 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)))).
18695 (* constant 3750 *)
18696 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)))))).
18697 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))).
18700 (* constant 3751 *)
18701 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)))).
18702 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)))).
18705 (* constant 3752 *)
18706 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)))).
18707 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)).
18710 (* constant 3753 *)
18711 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)).
18712 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)).
18715 (* constant 3754 *)
18716 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))).
18717 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))))).
18720 (* constant 3755 *)
18721 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))).
18722 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))))).
18725 (* constant 3756 *)
18726 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))).
18727 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)))).
18730 (* constant 3757 *)
18731 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)))))))))).
18732 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)))))))))).
18735 (* constant 3758 *)
18736 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)))))))))).
18737 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)))))))))).
18740 (* constant 3759 *)
18741 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)))))))))))).
18742 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)))))))))))).
18745 (* constant 3760 *)
18746 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)))))))))))))).
18747 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)))))))))))))).
18750 (* constant 3761 *)
18751 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)))))))))))))).
18752 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))))))))))))))).
18755 (* constant 3762 *)
18756 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)))))))))))))).
18757 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)))))))))))))))).
18760 (* constant 3763 *)
18761 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))))))))))).
18762 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)))))))))))))).
18765 (* constant 3764 *)
18766 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))))))))))).
18767 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))))))))))).
18770 (* constant 3765 *)
18771 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))))))))))).
18772 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)))))))))))).
18775 (* constant 3766 *)
18776 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)))))))))).
18777 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))))))))))).
18780 (* constant 3767 *)
18781 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))))))).
18782 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)))))))))).
18785 (* constant 3768 *)
18786 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))))))).
18787 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))))))))).
18790 (* constant 3769 *)
18791 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)))))).
18792 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))))))).
18795 (* constant 3770 *)
18796 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))))))))))).
18797 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)))))))))).
18800 (* constant 3771 *)
18801 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)))))))))).
18802 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))))))))))).
18805 (* constant 3772 *)
18806 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))))))))))).
18807 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))))))))))).
18810 (* constant 3773 *)
18811 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)))))))))))).
18812 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))))))))))))).
18815 (* constant 3774 *)
18816 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)))))))))))).
18817 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)))))))))))).
18820 (* constant 3775 *)
18821 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)))))))))))).
18822 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)))))))))))))).
18825 (* constant 3776 *)
18826 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)))))))))))).
18827 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))))))))))))).
18830 (* constant 3777 *)
18831 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))))))))))))).
18832 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)))))))))))).
18835 (* constant 3778 *)
18836 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))))))))))).
18837 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))))))))))).
18840 (* constant 3779 *)
18841 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)))))))))))).
18842 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))))))))))).
18845 (* constant 3780 *)
18846 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)))))))).
18847 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)))))))).
18850 (* constant 3781 *)
18851 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)))))))))))).
18852 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)))))))))))).
18855 (* constant 3782 *)
18856 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)))))))))))).
18857 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)))))))))))).
18860 (* constant 3783 *)
18861 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)))))))))))).
18862 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)))))))))))).
18865 (* constant 3784 *)
18866 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))))))))))))).
18867 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))))))))))))).
18870 (* constant 3785 *)
18871 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)))))))))))).
18872 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))))))))))))).
18875 (* constant 3786 *)
18876 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)))))))))).
18877 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)))))))))))).
18880 (* constant 3787 *)
18881 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))))))))))).
18882 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))))))))))).
18885 (* constant 3788 *)
18886 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)))))))))).
18887 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)))))))))).
18890 (* constant 3789 *)
18891 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)))))))))).
18892 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)))))))))).
18895 (* constant 3790 *)
18896 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)))))))))))))).
18897 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))))))))))))).
18900 (* constant 3791 *)
18901 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))))))))))))).
18902 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)))))))))))))).
18905 (* constant 3792 *)
18906 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)))))))))))))).
18907 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)))))))))))))).
18910 (* constant 3793 *)
18911 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))))))))))))).
18912 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))))))))))))).
18915 (* constant 3794 *)
18916 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)))))))))))))).
18917 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)))))))))))))).
18920 (* constant 3795 *)
18921 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))))))))))))))).
18922 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)))))))))))))).
18925 (* constant 3796 *)
18926 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)))))))))))).
18927 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))))))))))))).
18930 (* constant 3797 *)
18931 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)))))))))).
18932 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)))))))))).
18935 (* constant 3798 *)
18936 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))))))))))))))).
18937 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)))))))))))))))).
18940 (* constant 3799 *)
18941 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))))))))))))).
18942 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)))))))))))))).
18945 (* constant 3800 *)
18946 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))))))))))).
18947 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)))))))))))).
18950 (* constant 3801 *)
18951 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))))))))).
18952 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)))))))))).
18955 (* constant 3802 *)
18956 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))))))).
18957 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)))))))).
18960 (* constant 3803 *)
18961 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)))))).
18962 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))))))).
18965 (* constant 3804 *)
18966 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)))))))))))).
18967 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))))))))))).
18970 (* constant 3805 *)
18971 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))))))))))).
18972 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)))))))))))).
18975 (* constant 3806 *)
18976 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))))))))))))).
18977 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))))))))))))).
18980 (* constant 3807 *)
18981 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))))))))))))).
18982 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))))))))))))).
18985 (* constant 3808 *)
18986 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)))))))))))))).
18987 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)))))))))))))).
18990 (* constant 3809 *)
18991 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))))))))))))).
18992 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)))))))))))))).
18995 (* constant 3810 *)
18996 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)))))))))))))).
18997 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))))))))))))))).
19000 (* constant 3811 *)
19001 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)))))))))))))).
19002 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)))))))))))))).
19005 (* constant 3812 *)
19006 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))))))))))))).
19007 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)))))))))))))).
19010 (* constant 3813 *)
19011 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))))))))))).
19012 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))))))))))))).
19015 (* constant 3814 *)
19016 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)))))))).
19017 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))))))))))).
19020 (* constant 3815 *)
19021 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))))))))))))).
19022 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)))))))))))).
19025 (* constant 3816 *)
19026 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)))))))))))).
19027 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))))))))))))).
19030 (* constant 3817 *)
19031 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))))))))))))).
19032 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))))))))))))).
19035 (* constant 3818 *)
19036 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)))))))))))).
19037 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))))))))))))).
19040 (* constant 3819 *)
19041 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)))))))))))).
19042 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)))))))))))).
19045 (* constant 3820 *)
19046 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))))))))))).
19047 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)))))))))))).
19050 (* constant 3821 *)
19051 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)))))))).
19052 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))))))))))).
19055 (* constant 3822 *)
19056 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))))))).
19057 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)))))))).
19060 (* constant 3823 *)
19061 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))))))).
19062 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))))))).
19065 (* constant 3824 *)
19066 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)))))))))).
19067 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))))))).
19070 (* constant 3825 *)
19071 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)))))).
19072 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))))))).
19075 (* constant 3826 *)
19076 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)))))))).
19077 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)))))))).
19080 (* constant 3827 *)
19081 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)))))))).
19082 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)))))))).
19085 (* constant 3828 *)
19086 Definition l_e_st_eq_landau_n_rt_rp_ivad_i : l_e_st_eq_landau_n_rt_cut.
19087 exact l_e_st_eq_landau_n_rt_rp_1rp.
19090 (* constant 3829 *)
19091 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).
19092 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).
19095 (* constant 3830 *)
19096 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)).
19097 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)).
19100 (* constant 3831 *)
19101 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)).
19102 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)).
19105 (* constant 3832 *)
19106 Definition l_e_st_eq_landau_n_rt_rp_ivad_2 : l_e_st_eq_landau_n_rt_cut.
19107 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).
19110 (* constant 3833 *)
19111 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))).
19112 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)).
19115 (* constant 3834 *)
19116 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))).
19117 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))).
19120 (* constant 3835 *)
19121 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))).
19122 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))).
19125 (* constant 3836 *)
19126 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))).
19127 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))).
19130 (* constant 3837 *)
19131 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)))).
19132 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))).
19135 (* constant 3838 *)
19136 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)))).
19137 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))).
19140 (* constant 3839 *)
19141 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)).
19142 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)).
19145 (* constant 3840 *)
19146 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)))))).
19147 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)).
19150 (* constant 3841 *)
19151 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))).
19152 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)))).
19155 (* constant 3842 *)
19156 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)))).
19157 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))).
19160 (* constant 3843 *)
19161 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)))).
19162 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))).
19165 (* constant 3844 *)
19166 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)))).
19167 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))).
19170 (* constant 3845 *)
19171 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))))).
19172 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))))).
19175 (* constant 3846 *)
19176 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)))))).
19177 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)))).
19180 (* constant 3847 *)
19181 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)))).
19182 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))).
19185 (* constant 3848 *)
19186 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)))).
19187 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))).
19190 (* constant 3849 *)
19191 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)))).
19192 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))).
19195 (* constant 3850 *)
19196 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)))).
19197 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)))).
19200 (* constant 3851 *)
19201 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))).
19202 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)))).
19205 (* constant 3852 *)
19206 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)))))))).
19207 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))))))).
19210 (* constant 3853 *)
19211 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)))))))).
19212 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))))))))).
19215 (* constant 3854 *)
19216 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))))))))).
19217 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))))))))).
19220 (* constant 3855 *)
19221 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))))))))).
19222 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))))))))).
19225 (* constant 3856 *)
19226 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))))))))).
19227 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)))))))).
19230 (* constant 3857 *)
19231 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))))))))).
19232 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))))))))).
19235 (* constant 3858 *)
19236 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)))))))).
19237 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)))))))))).
19240 (* constant 3859 *)
19241 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)))))))).
19242 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))))))))).
19245 (* constant 3860 *)
19246 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)))))))).
19247 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)))))))).
19250 (* constant 3861 *)
19251 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))))))))).
19252 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))))))))).
19255 (* constant 3862 *)
19256 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)))))))).
19257 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))))))))).
19260 (* constant 3863 *)
19261 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)))))))))).
19262 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)))))))))).
19265 (* constant 3864 *)
19266 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))))))))).
19267 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))))))))).
19270 (* constant 3865 *)
19271 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)))))))).
19272 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))))))))).
19275 (* constant 3866 *)
19276 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))))))).
19277 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)))))))).
19280 (* constant 3867 *)
19281 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))).
19282 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))).
19285 (* constant 3868 *)
19286 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)))).
19287 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)))).
19290 (* constant 3869 *)
19291 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))))).
19292 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)))).
19295 (* constant 3870 *)
19296 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))).
19297 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))).
19300 (* constant 3871 *)
19301 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))).
19302 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)))).
19305 (* constant 3872 *)
19306 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))).
19307 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)))).
19310 (* constant 3873 *)
19311 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))).
19312 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)))).
19315 (* constant 3874 *)
19316 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))).
19317 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))))).
19320 (* constant 3875 *)
19321 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)))).
19322 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)))).
19325 (* constant 3876 *)
19326 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))))).
19327 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)))).
19330 (* constant 3877 *)
19331 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)))).
19332 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))).
19335 (* constant 3878 *)
19336 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)))).
19337 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)))).
19340 (* constant 3879 *)
19341 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)))).
19342 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))).
19345 (* constant 3880 *)
19346 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))).
19347 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))).
19350 (* constant 3881 *)
19351 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))))).
19352 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))).
19355 (* constant 3882 *)
19356 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))).
19357 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))).
19360 (* constant 3883 *)
19361 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)))))).
19362 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)))))).
19365 (* constant 3884 *)
19366 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)))))).
19367 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))))).
19370 (* constant 3885 *)
19371 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)))))).
19372 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)))))).
19375 (* constant 3886 *)
19376 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)))))).
19377 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))))).
19380 (* constant 3887 *)
19381 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)))).
19382 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))).
19385 (* constant 3888 *)
19386 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))))).
19387 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))))).
19390 (* constant 3889 *)
19391 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)))))).
19392 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)))))).
19395 (* constant 3890 *)
19396 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))))))).
19397 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))))))).
19400 (* constant 3891 *)
19401 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))))))))).
19402 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)))))))).
19405 (* constant 3892 *)
19406 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)))))))))).
19407 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)))))))).
19410 (* constant 3893 *)
19411 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)))))))))).
19412 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))))))))).
19415 (* constant 3894 *)
19416 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))))))))).
19417 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)))))))))).
19420 (* constant 3895 *)
19421 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)))))))).
19422 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)))))))).
19425 (* constant 3896 *)
19426 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)))))))).
19427 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)))))))).
19430 (* constant 3897 *)
19431 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)))))))).
19432 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)))))))).
19435 (* constant 3898 *)
19436 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)))))))).
19437 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))))))).
19440 (* constant 3899 *)
19441 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))))))).
19442 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)))))))).
19445 (* constant 3900 *)
19446 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))))))))).
19447 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)))))))).
19450 (* constant 3901 *)
19451 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))))))))).
19452 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))))))).
19455 (* constant 3902 *)
19456 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)))))))).
19457 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)))))))).
19460 (* constant 3903 *)
19461 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))))))))).
19462 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)))))))).
19465 (* constant 3904 *)
19466 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)))))))).
19467 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)))))))).
19470 (* constant 3905 *)
19471 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))))))))).
19472 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)))))))).
19475 (* constant 3906 *)
19476 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)))))))))).
19477 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)))))))).
19480 (* constant 3907 *)
19481 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)))))))))).
19482 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))))))))).
19485 (* constant 3908 *)
19486 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)))))))).
19487 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))))))))).
19490 (* constant 3909 *)
19491 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))))))))).
19492 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)))))))).
19495 (* constant 3910 *)
19496 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)))))))).
19497 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)))))))).
19500 (* constant 3911 *)
19501 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)))))))).
19502 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)))))))).
19505 (* constant 3912 *)
19506 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)))))))).
19507 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)))))))).
19510 (* constant 3913 *)
19511 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))))))).
19512 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))))))).
19515 (* constant 3914 *)
19516 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))))))).
19517 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))))))).
19520 (* constant 3915 *)
19521 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))))))).
19522 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))))))).
19525 (* constant 3916 *)
19526 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)))))).
19527 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))))))).
19530 (* constant 3917 *)
19531 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))))))).
19532 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)))))).
19535 (* constant 3918 *)
19536 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))))))).
19537 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)))))).
19540 (* constant 3919 *)
19541 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))))))).
19542 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))))))).
19545 (* constant 3920 *)
19546 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)))))).
19547 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)))))).
19550 (* constant 3921 *)
19551 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)))))).
19552 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))))).
19555 (* constant 3922 *)
19556 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))))))).
19557 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))))).
19560 (* constant 3923 *)
19561 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))))))).
19562 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)))))).
19565 (* constant 3924 *)
19566 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))))))).
19567 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)))))).
19570 (* constant 3925 *)
19571 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))))))).
19572 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)))))).
19575 (* constant 3926 *)
19576 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)))))))).
19577 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)))))).
19580 (* constant 3927 *)
19581 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)))))).
19582 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)))))).
19585 (* constant 3928 *)
19586 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))))).
19587 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))))).
19590 (* constant 3929 *)
19591 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))))).
19592 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))))).
19595 (* constant 3930 *)
19596 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)))))).
19597 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)))))).
19600 (* constant 3931 *)
19601 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)))))).
19602 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)))))).
19605 (* constant 3932 *)
19606 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)))))).
19607 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)))))).
19610 (* constant 3933 *)
19611 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))))).
19612 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)))))).
19615 (* constant 3934 *)
19616 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)))))).
19617 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)))))).
19620 (* constant 3935 *)
19621 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))))).
19622 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)))))).
19625 (* constant 3936 *)
19626 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))))))).
19627 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))))))).
19630 (* constant 3937 *)
19631 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))))))).
19632 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))))).
19635 (* constant 3938 *)
19636 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)))))).
19637 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)))))).
19640 (* constant 3939 *)
19641 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))))))))).
19642 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)))))))).
19645 (* constant 3940 *)
19646 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))))))))).
19647 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)))))))).
19650 (* constant 3941 *)
19651 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)))))))).
19652 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))))))))).
19655 (* constant 3942 *)
19656 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))))))).
19657 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))))))).
19660 (* constant 3943 *)
19661 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))))))).
19662 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)))))).
19665 (* constant 3944 *)
19666 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))))).
19667 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))))).
19670 (* constant 3945 *)
19671 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)).
19672 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)).
19675 (* constant 3946 *)
19676 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).
19677 exact (fun (x:l_e_st_eq_landau_n_rt_rp_dif) => l_e_st_eq_landau_n_rt_rp_refeq x).
19680 (* constant 3947 *)
19681 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))).
19682 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))).
19685 (* constant 3948 *)
19686 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))))).
19687 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))))).
19690 (* constant 3949 *)
19691 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)).
19692 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)).
19695 (* constant 3950 *)
19696 Definition l_e_st_eq_landau_n_rt_rp_r_real : Type.
19697 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).
19700 (* constant 3951 *)
19701 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)).
19702 exact (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)).
19705 (* constant 3952 *)
19706 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)).
19707 exact (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))).
19710 (* constant 3953 *)
19711 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).
19712 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).
19715 (* constant 3954 *)
19716 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).
19717 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).
19720 (* constant 3955 *)
19721 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).
19722 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).
19725 (* constant 3956 *)
19726 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)).
19727 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)).
19730 (* constant 3957 *)
19731 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).
19732 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).
19735 (* constant 3958 *)
19736 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).
19737 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).
19740 (* constant 3959 *)
19741 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))).
19742 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).
19745 (* constant 3960 *)
19746 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)))))).
19747 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))))).
19750 (* constant 3961 *)
19751 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))).
19752 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))).
19755 (* constant 3962 *)
19756 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)))))).
19757 exact (fun (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))))))).
19760 (* constant 3963 *)
19761 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)))).
19762 exact (fun (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)))))).
19765 (* constant 3964 *)
19766 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))))))).
19767 exact (fun (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))))))))).
19770 (* constant 3965 *)
19771 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))))).
19772 exact (fun (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))))))).
19775 (* constant 3966 *)
19776 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)))))))).
19777 exact (fun (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))))))))))).
19780 (* constant 3967 *)
19781 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)))))).
19782 exact (fun (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)))))))).
19785 (* constant 3968 *)
19786 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))))))).
19787 exact (fun (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))))))).
19790 (* constant 3969 *)
19791 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))))))).
19792 exact (fun (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))))))).
19795 (* constant 3970 *)
19796 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))))))).
19797 exact (fun (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)))))))).
19800 (* constant 3971 *)
19801 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)))))))).
19802 exact (fun (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)))))))).
19805 (* constant 3972 *)
19806 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)).
19807 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)).
19810 (* constant 3973 *)
19811 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)))).
19812 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)))).
19815 (* constant 3974 *)
19816 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))))))).
19817 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)))))).
19820 (* constant 3975 *)
19821 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)).
19822 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)).
19825 (* constant 3976 *)
19826 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))))).
19827 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))))).
19830 (* constant 3977 *)
19831 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)))))))))).
19832 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))))))))).
19835 (* constant 3978 *)
19836 Definition l_e_st_eq_landau_n_rt_rp_r_0 : l_e_st_eq_landau_n_rt_rp_r_real.
19837 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)).
19840 (* constant 3979 *)
19841 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)))).
19842 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))))))).
19845 (* constant 3980 *)
19846 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)))).
19847 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)))))).
19850 (* constant 3981 *)
19851 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)).
19852 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))).
19855 (* constant 3982 *)
19856 Definition l_e_st_eq_landau_n_rt_rp_r_pos : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop).
19857 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)).
19860 (* constant 3983 *)
19861 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)))).
19862 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)))).
19865 (* constant 3984 *)
19866 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)))).
19867 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))))).
19870 (* constant 3985 *)
19871 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))))))).
19872 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)))))).
19875 (* constant 3986 *)
19876 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)))))).
19877 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)))))).
19880 (* constant 3987 *)
19881 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)))))).
19882 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))))))).
19885 (* constant 3988 *)
19886 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)))).
19887 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)))))).
19890 (* constant 3989 *)
19891 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)).
19892 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))).
19895 (* constant 3990 *)
19896 Definition l_e_st_eq_landau_n_rt_rp_r_neg : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop).
19897 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)).
19900 (* constant 3991 *)
19901 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)))).
19902 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)))).
19905 (* constant 3992 *)
19906 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)))).
19907 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))))).
19910 (* constant 3993 *)
19911 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))))))).
19912 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)))))).
19915 (* constant 3994 *)
19916 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)))))).
19917 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)))))).
19920 (* constant 3995 *)
19921 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)))))).
19922 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))))))).
19925 (* constant 3996 *)
19926 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)))).
19927 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)))))).
19930 (* constant 3997 *)
19931 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))))).
19932 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))))).
19935 (* constant 3998 *)
19936 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))))).
19937 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))))).
19940 (* constant 3999 *)
19941 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))))).
19942 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))))).
19945 (* constant 4000 *)
19946 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)))).
19947 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)))).
19950 (* constant 4001 *)
19951 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)).
19952 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))).
19955 (* constant 4002 *)
19956 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))))).
19957 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))))).
19960 (* constant 4003 *)
19961 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))))).
19962 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))))).
19965 (* constant 4004 *)
19966 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))))).
19967 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))))).
19970 (* constant 4005 *)
19971 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)))).
19972 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)))).
19975 (* constant 4006 *)
19976 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)).
19977 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))).
19980 (* constant 4007 *)
19981 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)).
19982 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)).
19985 (* constant 4008 *)
19986 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))))).
19987 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))))).
19990 (* constant 4009 *)
19991 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))).
19992 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)).
19995 (* constant 4010 *)
19996 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)).
19997 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)).
20000 (* constant 4011 *)
20001 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))).
20002 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)).
20005 (* constant 4012 *)
20006 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))).
20007 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)).
20010 (* constant 4013 *)
20011 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))).
20012 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)).
20015 (* constant 4014 *)
20016 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)).
20017 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)).
20020 (* constant 4015 *)
20021 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)))).
20022 exact (fun (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)))).
20025 (* constant 4016 *)
20026 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)))).
20027 exact (fun (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)))).
20030 (* constant 4017 *)
20031 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).
20032 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)).
20035 (* constant 4018 *)
20036 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).
20037 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)).
20040 (* constant 4019 *)
20041 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)))).
20042 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))).
20045 (* constant 4020 *)
20046 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)))).
20047 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))).
20050 (* constant 4021 *)
20051 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)))).
20052 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))).
20055 (* constant 4022 *)
20056 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))).
20057 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)))).
20060 (* constant 4023 *)
20061 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)))).
20062 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))).
20065 (* constant 4024 *)
20066 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))).
20067 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)))).
20070 (* constant 4025 *)
20071 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)).
20072 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)).
20075 (* constant 4026 *)
20076 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)).
20077 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)).
20080 (* constant 4027 *)
20081 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))))))).
20082 exact (fun (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)))))))).
20085 (* constant 4028 *)
20086 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))).
20087 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)))))).
20090 (* constant 4029 *)
20091 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)))).
20092 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)))).
20095 (* constant 4030 *)
20096 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)))).
20097 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))))).
20100 (* constant 4031 *)
20101 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)))))).
20102 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)))))).
20105 (* constant 4032 *)
20106 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)))))).
20107 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))))).
20110 (* constant 4033 *)
20111 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)))).
20112 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)))).
20115 (* constant 4034 *)
20116 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)))).
20117 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))).
20120 (* constant 4035 *)
20121 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)).
20122 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))).
20125 (* constant 4036 *)
20126 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)))).
20127 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))).
20130 (* constant 4037 *)
20131 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)).
20132 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))).
20135 (* constant 4038 *)
20136 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)))))).
20137 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))))).
20140 (* constant 4039 *)
20141 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))))).
20142 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)))))).
20145 (* constant 4040 *)
20146 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))).
20147 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))).
20150 (* constant 4041 *)
20151 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).
20152 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)).
20155 (* constant 4042 *)
20156 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))))))).
20157 exact (fun (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)))))))).
20160 (* constant 4043 *)
20161 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))).
20162 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)))))).
20165 (* constant 4044 *)
20166 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)))).
20167 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)))).
20170 (* constant 4045 *)
20171 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)))).
20172 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))))).
20175 (* constant 4046 *)
20176 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)))))).
20177 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)))))).
20180 (* constant 4047 *)
20181 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)))))).
20182 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))))).
20185 (* constant 4048 *)
20186 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)))).
20187 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)))).
20190 (* constant 4049 *)
20191 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)))).
20192 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))).
20195 (* constant 4050 *)
20196 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)).
20197 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))).
20200 (* constant 4051 *)
20201 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)))).
20202 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))).
20205 (* constant 4052 *)
20206 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)).
20207 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))).
20210 (* constant 4053 *)
20211 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)))))).
20212 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))))).
20215 (* constant 4054 *)
20216 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))))).
20217 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)))))).
20220 (* constant 4055 *)
20221 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))).
20222 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))).
20225 (* constant 4056 *)
20226 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).
20227 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)).
20230 (* constant 4057 *)
20231 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).
20232 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).
20235 (* constant 4058 *)
20236 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))).
20237 exact (fun (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))).
20240 (* constant 4059 *)
20241 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))))).
20242 exact (fun (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))))).
20245 (* constant 4060 *)
20246 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).
20247 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)).
20250 (* constant 4061 *)
20251 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)))).
20252 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)))).
20255 (* constant 4062 *)
20256 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.
20257 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))).
20260 (* constant 4063 *)
20261 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).
20262 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).
20265 (* constant 4064 *)
20266 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)))).
20267 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))).
20270 (* constant 4065 *)
20271 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))))).
20272 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)))).
20275 (* constant 4066 *)
20276 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)))).
20277 exact (fun (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))).
20280 (* constant 4067 *)
20281 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))))).
20282 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))))).
20285 (* constant 4068 *)
20286 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))))).
20287 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))))).
20290 (* constant 4069 *)
20291 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))).
20292 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)))).
20295 (* constant 4070 *)
20296 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))))).
20297 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))))).
20300 (* constant 4071 *)
20301 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))))).
20302 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))))).
20305 (* constant 4072 *)
20306 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))).
20307 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)))).
20310 (* constant 4073 *)
20311 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))))))))).
20312 exact (fun (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)))))))))).
20315 (* constant 4074 *)
20316 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))))))))).
20317 exact (fun (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)))))))))).
20320 (* constant 4075 *)
20321 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))))).
20322 exact (fun (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))))))))).
20325 (* constant 4076 *)
20326 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))))))))).
20327 exact (fun (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)))))))))).
20330 (* constant 4077 *)
20331 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))))))))).
20332 exact (fun (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)))))))))).
20335 (* constant 4078 *)
20336 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))))).
20337 exact (fun (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))))))))).
20340 (* constant 4079 *)
20341 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))).
20342 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))).
20345 (* constant 4080 *)
20346 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))))).
20347 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))))).
20350 (* constant 4081 *)
20351 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)))).
20352 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))))).
20355 (* constant 4082 *)
20356 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)).
20357 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)))).
20360 (* constant 4083 *)
20361 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)).
20362 exact (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))))).
20365 (* constant 4084 *)
20366 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)))).
20367 exact (fun (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))))).
20370 (* constant 4085 *)
20371 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))))))).
20372 exact (fun (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))))))).
20375 (* constant 4086 *)
20376 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)))))))).
20377 exact (fun (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)))))))).
20380 (* constant 4087 *)
20381 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))))))).
20382 exact (fun (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)))))))).
20385 (* constant 4088 *)
20386 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)))))))))))).
20387 exact (fun (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))))))))))).
20390 (* constant 4089 *)
20391 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)))))))))))).
20392 exact (fun (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))))))))))).
20395 (* constant 4090 *)
20396 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))))))))))).
20397 exact (fun (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))))))))))).
20400 (* constant 4091 *)
20401 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))))))))))).
20402 exact (fun (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)))))))))))).
20405 (* constant 4092 *)
20406 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))))))))).
20407 exact (fun (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))))))))))).
20410 (* constant 4093 *)
20411 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))))))).
20412 exact (fun (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))))))))).
20415 (* constant 4094 *)
20416 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)).
20417 exact (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))))).
20420 (* constant 4095 *)
20421 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)))).
20422 exact (fun (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))))).
20425 (* constant 4096 *)
20426 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))))))).
20427 exact (fun (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))))))).
20430 (* constant 4097 *)
20431 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)))))))).
20432 exact (fun (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)))))))).
20435 (* constant 4098 *)
20436 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))))))).
20437 exact (fun (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)))))))).
20440 (* constant 4099 *)
20441 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)))))))))))).
20442 exact (fun (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))))))))))).
20445 (* constant 4100 *)
20446 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)))))))))))).
20447 exact (fun (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))))))))))).
20450 (* constant 4101 *)
20451 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))))))))))).
20452 exact (fun (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))))))))))).
20455 (* constant 4102 *)
20456 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))))))))))).
20457 exact (fun (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)))))))))))).
20460 (* constant 4103 *)
20461 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))))))))).
20462 exact (fun (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))))))))))).
20465 (* constant 4104 *)
20466 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))))))).
20467 exact (fun (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))))))))).
20470 (* constant 4105 *)
20471 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))))).
20472 exact (fun (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))))).
20475 (* constant 4106 *)
20476 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))))).
20477 exact (fun (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))))).
20480 (* constant 4107 *)
20481 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))))).
20482 exact (fun (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))))).
20485 (* constant 4108 *)
20486 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))))).
20487 exact (fun (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))))).
20490 (* constant 4109 *)
20491 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))))))).
20492 exact (fun (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)))))))).
20495 (* constant 4110 *)
20496 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))))))).
20497 exact (fun (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)))))))).
20500 (* constant 4111 *)
20501 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))))))).
20502 exact (fun (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)))))))).
20505 (* constant 4112 *)
20506 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))))))).
20507 exact (fun (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)))))))).
20510 (* constant 4113 *)
20511 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))).
20512 exact (fun (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))))))).
20515 (* constant 4114 *)
20516 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))))))).
20517 exact (fun (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)))))))).
20520 (* constant 4115 *)
20521 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))))))).
20522 exact (fun (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)))))))).
20525 (* constant 4116 *)
20526 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))).
20527 exact (fun (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))))))).
20530 (* constant 4117 *)
20531 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))))))).
20532 exact (fun (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)))))).
20535 (* constant 4118 *)
20536 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)))))))).
20537 exact (fun (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)))))))).
20540 (* constant 4119 *)
20541 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)))))))).
20542 exact (fun (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)))))))).
20545 (* constant 4120 *)
20546 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)))))))).
20547 exact (fun (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)))))))).
20550 (* constant 4121 *)
20551 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))))))).
20552 exact (fun (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))))))).
20555 (* constant 4122 *)
20556 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))))))).
20557 exact (fun (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)))))).
20560 (* constant 4123 *)
20561 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)))))))).
20562 exact (fun (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)))))))).
20565 (* constant 4124 *)
20566 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)))))))).
20567 exact (fun (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)))))))).
20570 (* constant 4125 *)
20571 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)))))))).
20572 exact (fun (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)))))))).
20575 (* constant 4126 *)
20576 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))))))).
20577 exact (fun (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)))))))).
20580 (* constant 4127 *)
20581 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))))))).
20582 exact (fun (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))))))).
20585 (* constant 4128 *)
20586 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))).
20587 exact (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)))))).
20590 (* constant 4129 *)
20591 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))).
20592 exact (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))).
20595 (* constant 4130 *)
20596 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))).
20597 exact (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))).
20600 (* constant 4131 *)
20601 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)).
20602 exact (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))).
20605 (* constant 4132 *)
20606 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)).
20607 exact (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))).
20610 (* constant 4133 *)
20611 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))).
20612 exact (fun (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)))).
20615 (* constant 4134 *)
20616 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))).
20617 exact (fun (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)))).
20620 (* constant 4135 *)
20621 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))))).
20622 exact (fun (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))))).
20625 (* constant 4136 *)
20626 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))))).
20627 exact (fun (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))))).
20630 (* constant 4137 *)
20631 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))))).
20632 exact (fun (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))))).
20635 (* constant 4138 *)
20636 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))))).
20637 exact (fun (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))))).
20640 (* constant 4139 *)
20641 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))))))).
20642 exact (fun (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)))))))).
20645 (* constant 4140 *)
20646 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))))))).
20647 exact (fun (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)))))))).
20650 (* constant 4141 *)
20651 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))).
20652 exact (fun (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))).
20655 (* constant 4142 *)
20656 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))).
20657 exact (fun (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))).
20660 (* constant 4143 *)
20661 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))).
20662 exact (fun (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))).
20665 (* constant 4144 *)
20666 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))).
20667 exact (fun (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))).
20670 (* constant 4145 *)
20671 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))))))).
20672 exact (fun (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))))))))).
20675 (* constant 4146 *)
20676 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))))))).
20677 exact (fun (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))))))))).
20680 (* constant 4147 *)
20681 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))))))).
20682 exact (fun (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))))))))).
20685 (* constant 4148 *)
20686 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))))))).
20687 exact (fun (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))))))))).
20690 (* constant 4149 *)
20691 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)))).
20692 exact (fun (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)))).
20695 (* constant 4150 *)
20696 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)))).
20697 exact (fun (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))).
20700 (* constant 4151 *)
20701 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))).
20702 exact (fun (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))).
20705 (* constant 4152 *)
20706 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))).
20707 exact (fun (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)))).
20710 (* constant 4153 *)
20711 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)))).
20712 exact (fun (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)))).
20715 (* constant 4154 *)
20716 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)))).
20717 exact (fun (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)))).
20720 (* constant 4155 *)
20721 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))).
20722 exact (fun (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)))).
20725 (* constant 4156 *)
20726 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))).
20727 exact (fun (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)))).
20730 (* constant 4157 *)
20731 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)))))).
20732 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))))))).
20735 (* constant 4158 *)
20736 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)))))).
20737 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))))))).
20740 (* constant 4159 *)
20741 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)).
20742 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)))))).
20745 (* constant 4160 *)
20746 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)))))).
20747 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))))))).
20750 (* constant 4161 *)
20751 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)))))).
20752 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))))))).
20755 (* constant 4162 *)
20756 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)).
20757 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)))))).
20760 (* constant 4163 *)
20761 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)))))).
20762 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))))))).
20765 (* constant 4164 *)
20766 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)))))).
20767 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))))))).
20770 (* constant 4165 *)
20771 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)).
20772 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)))))).
20775 (* constant 4166 *)
20776 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)))))).
20777 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))))))).
20780 (* constant 4167 *)
20781 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)))))).
20782 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))))))).
20785 (* constant 4168 *)
20786 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)).
20787 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)))))).
20790 (* constant 4169 *)
20791 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))))).
20792 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))))))).
20795 (* constant 4170 *)
20796 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))))).
20797 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)))))).
20800 (* constant 4171 *)
20801 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).
20802 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))))).
20805 (* constant 4172 *)
20806 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))).
20807 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)).
20810 (* constant 4173 *)
20811 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))))))))))).
20812 exact (fun (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)))))))))))).
20815 (* constant 4174 *)
20816 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))))))))))).
20817 exact (fun (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)))))))))))).
20820 (* constant 4175 *)
20821 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))))).
20822 exact (fun (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))))))))))).
20825 (* constant 4176 *)
20826 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))))).
20827 exact (fun (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))))).
20830 (* constant 4177 *)
20831 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))))).
20832 exact (fun (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))))))).
20835 (* constant 4178 *)
20836 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))))))))))).
20837 exact (fun (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)))))))))))).
20840 (* constant 4179 *)
20841 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))))))))))).
20842 exact (fun (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)))))))))))).
20845 (* constant 4180 *)
20846 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))))).
20847 exact (fun (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))))))))))).
20850 (* constant 4181 *)
20851 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))))))))))).
20852 exact (fun (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)))))))))))).
20855 (* constant 4182 *)
20856 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))))))))))).
20857 exact (fun (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)))))))))))).
20860 (* constant 4183 *)
20861 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))))).
20862 exact (fun (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))))))))))).
20865 (* constant 4184 *)
20866 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))))).
20867 exact (fun (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))))))).
20870 (* constant 4185 *)
20871 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))))).
20872 exact (fun (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))))))).
20875 (* constant 4186 *)
20876 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))))))))))).
20877 exact (fun (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)))))))))))).
20880 (* constant 4187 *)
20881 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))))))))))).
20882 exact (fun (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)))))))))))).
20885 (* constant 4188 *)
20886 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))))).
20887 exact (fun (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))))))))))).
20890 (* constant 4189 *)
20891 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))))).
20892 exact (fun (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))))).
20895 (* constant 4190 *)
20896 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))))).
20897 exact (fun (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))))))).
20900 (* constant 4191 *)
20901 Definition l_e_st_eq_landau_n_rt_rp_r_ratrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop).
20902 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))).
20905 (* constant 4192 *)
20906 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))))).
20907 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)))).
20910 (* constant 4193 *)
20911 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)))).
20912 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))))).
20915 (* constant 4194 *)
20916 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))))))).
20917 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)))))).
20920 (* constant 4195 *)
20921 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)))))).
20922 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)))))).
20925 (* constant 4196 *)
20926 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)))))).
20927 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))))))).
20930 (* constant 4197 *)
20931 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)))).
20932 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)))))).
20935 (* constant 4198 *)
20936 Definition l_e_st_eq_landau_n_rt_rp_r_irratrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop).
20937 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)).
20940 (* constant 4199 *)
20941 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))).
20942 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))).
20945 (* constant 4200 *)
20946 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))).
20947 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))).
20950 (* constant 4201 *)
20951 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))).
20952 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))).
20955 (* constant 4202 *)
20956 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))).
20957 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))).
20960 (* constant 4203 *)
20961 Definition l_e_st_eq_landau_n_rt_rp_r_natrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop).
20962 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))).
20965 (* constant 4204 *)
20966 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))))).
20967 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)))).
20970 (* constant 4205 *)
20971 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)))).
20972 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))))).
20975 (* constant 4206 *)
20976 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))))))).
20977 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)))))).
20980 (* constant 4207 *)
20981 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)))))).
20982 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)))))).
20985 (* constant 4208 *)
20986 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)))))).
20987 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))))))).
20990 (* constant 4209 *)
20991 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)))).
20992 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)))))).
20995 (* constant 4210 *)
20996 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)))).
20997 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))))).
21000 (* constant 4211 *)
21001 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)))).
21002 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))))).
21005 (* constant 4212 *)
21006 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)).
21007 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)))).
21010 (* constant 4213 *)
21011 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).
21012 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)).
21015 (* constant 4214 *)
21016 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)).
21017 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)).
21020 (* constant 4215 *)
21021 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)))).
21022 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))).
21025 (* constant 4216 *)
21026 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))).
21027 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)))).
21030 (* constant 4217 *)
21031 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).
21032 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))).
21035 (* constant 4218 *)
21036 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)))).
21037 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))))).
21040 (* constant 4219 *)
21041 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)))).
21042 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))))).
21045 (* constant 4220 *)
21046 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))))).
21047 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))))).
21050 (* constant 4221 *)
21051 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)))).
21052 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))))).
21055 (* constant 4222 *)
21056 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)))))).
21057 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)))))).
21060 (* constant 4223 *)
21061 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)))))).
21062 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))))).
21065 (* constant 4224 *)
21066 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)))))).
21067 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))))).
21070 (* constant 4225 *)
21071 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)))).
21072 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))))).
21075 (* constant 4226 *)
21076 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)).
21077 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)))).
21080 (* constant 4227 *)
21081 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)))).
21082 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)))).
21085 (* constant 4228 *)
21086 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)).
21087 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)))).
21090 (* constant 4229 *)
21091 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)).
21092 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))).
21095 (* constant 4230 *)
21096 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)))))).
21097 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))))).
21100 (* constant 4231 *)
21101 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))))).
21102 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))))).
21105 (* constant 4232 *)
21106 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)))).
21107 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))).
21110 (* constant 4233 *)
21111 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)).
21112 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))).
21115 (* constant 4234 *)
21116 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).
21117 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)).
21120 (* constant 4235 *)
21121 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))).
21122 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))).
21125 (* constant 4236 *)
21126 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))).
21127 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)).
21130 (* constant 4237 *)
21131 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).
21132 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)).
21135 (* constant 4238 *)
21136 Definition l_e_st_eq_landau_n_rt_rp_r_intrl : (forall (r:l_e_st_eq_landau_n_rt_rp_r_real), Prop).
21137 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))).
21140 (* constant 4239 *)
21141 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))))).
21142 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)))).
21145 (* constant 4240 *)
21146 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)))).
21147 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))))).
21150 (* constant 4241 *)
21151 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))))))).
21152 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)))))).
21155 (* constant 4242 *)
21156 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)))))).
21157 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)))))).
21160 (* constant 4243 *)
21161 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)))))).
21162 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))))))).
21165 (* constant 4244 *)
21166 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)))).
21167 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)))))).
21170 (* constant 4245 *)
21171 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)))).
21172 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))))).
21175 (* constant 4246 *)
21176 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)))).
21177 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))))).
21180 (* constant 4247 *)
21181 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)).
21182 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)))).
21185 (* constant 4248 *)
21186 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))))).
21187 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)))))).
21190 (* constant 4249 *)
21191 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))))).
21192 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)))))).
21195 (* constant 4250 *)
21196 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))).
21197 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))))).
21200 (* constant 4251 *)
21201 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)))).
21202 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))))).
21205 (* constant 4252 *)
21206 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)))).
21207 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))))).
21210 (* constant 4253 *)
21211 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)).
21212 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)))).
21215 (* constant 4254 *)
21216 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))).
21217 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))).
21220 (* constant 4255 *)
21221 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))).
21222 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))).
21225 (* constant 4256 *)
21226 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)))).
21227 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))))).
21230 (* constant 4257 *)
21231 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)))).
21232 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))))).
21235 (* constant 4258 *)
21236 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)).
21237 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)))).
21240 (* constant 4259 *)
21241 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)).
21242 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))).
21245 (* constant 4260 *)
21246 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))))))).
21247 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))))))).
21250 (* constant 4261 *)
21251 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.
21252 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)))))).
21255 (* constant 4262 *)
21256 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)).
21257 exact (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)).
21260 (* constant 4263 *)
21261 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))))))).
21262 exact (fun (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)))))).
21265 (* constant 4264 *)
21266 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)))))))).
21267 exact (fun (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))))))).
21270 (* constant 4265 *)
21271 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))))).
21272 exact (fun (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)))).
21275 (* constant 4266 *)
21276 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))))).
21277 exact (fun (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)))).
21280 (* constant 4267 *)
21281 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))))))).
21282 exact (fun (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))))))).
21285 (* constant 4268 *)
21286 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))))))).
21287 exact (fun (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)))))).
21290 (* constant 4269 *)
21291 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))))))).
21292 exact (fun (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))))))).
21295 (* constant 4270 *)
21296 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))).
21297 exact (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)))))).
21300 (* constant 4271 *)
21301 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))).
21302 exact (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)).
21305 (* constant 4272 *)
21306 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))))))).
21307 exact (fun (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)))))))).
21310 (* constant 4273 *)
21311 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))))))).
21312 exact (fun (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)))))))).
21315 (* constant 4274 *)
21316 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))).
21317 exact (fun (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))))))).
21320 (* constant 4275 *)
21321 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))).
21322 exact (fun (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)))).
21325 (* constant 4276 *)
21326 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))))))))).
21327 exact (fun (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))))))))).
21330 (* constant 4277 *)
21331 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))))))))).
21332 exact (fun (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))))))))).
21335 (* constant 4278 *)
21336 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))))).
21337 exact (fun (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)))))))).
21340 (* constant 4279 *)
21341 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))))))))).
21342 exact (fun (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))))))))).
21345 (* constant 4280 *)
21346 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))))))))).
21347 exact (fun (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))))))))).
21350 (* constant 4281 *)
21351 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))))).
21352 exact (fun (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)))))))).
21355 (* constant 4282 *)
21356 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).
21357 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)).
21360 (* constant 4283 *)
21361 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)))).
21362 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)))).
21365 (* constant 4284 *)
21366 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.
21367 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))).
21370 (* constant 4285 *)
21371 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).
21372 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).
21375 (* constant 4286 *)
21376 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)))).
21377 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))).
21380 (* constant 4287 *)
21381 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))))).
21382 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)))).
21385 (* constant 4288 *)
21386 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)))).
21387 exact (fun (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))).
21390 (* constant 4289 *)
21391 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))))).
21392 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))))).
21395 (* constant 4290 *)
21396 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))))).
21397 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))))).
21400 (* constant 4291 *)
21401 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))).
21402 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)))).
21405 (* constant 4292 *)
21406 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)))).
21407 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)))))).
21410 (* constant 4293 *)
21411 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)))).
21412 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))))).
21415 (* constant 4294 *)
21416 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)).
21417 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)))).
21420 (* constant 4295 *)
21421 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)).
21422 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))).
21425 (* constant 4296 *)
21426 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)).
21427 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)).
21430 (* constant 4297 *)
21431 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))))).
21432 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))))).
21435 (* constant 4298 *)
21436 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))))).
21437 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))))).
21440 (* constant 4299 *)
21441 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))).
21442 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)))).
21445 (* constant 4300 *)
21446 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))))).
21447 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))))).
21450 (* constant 4301 *)
21451 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)))).
21452 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))))).
21455 (* constant 4302 *)
21456 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)).
21457 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)))).
21460 (* constant 4303 *)
21461 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))))).
21462 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))))).
21465 (* constant 4304 *)
21466 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))))).
21467 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))))).
21470 (* constant 4305 *)
21471 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))).
21472 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)))).
21475 (* constant 4306 *)
21476 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)))).
21477 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))))).
21480 (* constant 4307 *)
21481 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)))).
21482 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))))).
21485 (* constant 4308 *)
21486 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)).
21487 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)))).
21490 (* constant 4309 *)
21491 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)))).
21492 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))))).
21495 (* constant 4310 *)
21496 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)))).
21497 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))))).
21500 (* constant 4311 *)
21501 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)).
21502 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)))).
21505 (* constant 4312 *)
21506 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)))).
21507 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))))).
21510 (* constant 4313 *)
21511 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)))).
21512 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))))).
21515 (* constant 4314 *)
21516 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)).
21517 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)))).
21520 (* constant 4315 *)
21521 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))).
21522 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)))).
21525 (* constant 4316 *)
21526 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).
21527 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))).
21530 (* constant 4317 *)
21531 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))).
21532 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)).
21535 (* constant 4318 *)
21536 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))).
21537 exact (fun (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)))).
21540 (* constant 4319 *)
21541 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)))).
21542 exact (fun (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)))).
21545 (* constant 4320 *)
21546 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)))).
21547 exact (fun (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)))).
21550 (* constant 4321 *)
21551 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))).
21552 exact (fun (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)))).
21555 (* constant 4322 *)
21556 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)))).
21557 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)))).
21560 (* constant 4323 *)
21561 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)).
21562 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))).
21565 (* constant 4324 *)
21566 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))).
21567 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)).
21570 (* constant 4325 *)
21571 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))).
21572 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)))).
21575 (* constant 4326 *)
21576 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).
21577 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))).
21580 (* constant 4327 *)
21581 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).
21582 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)).
21585 (* constant 4328 *)
21586 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)))))))).
21587 exact (fun (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))))))).
21590 (* constant 4329 *)
21591 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)))).
21592 exact (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)))))).
21595 (* constant 4330 *)
21596 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)))).
21597 exact (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))).
21600 (* constant 4331 *)
21601 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)).
21602 exact (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))).
21605 (* constant 4332 *)
21606 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)))))))).
21607 exact (fun (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))))))).
21610 (* constant 4333 *)
21611 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))))).
21612 exact (fun (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)))).
21615 (* constant 4334 *)
21616 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))))).
21617 exact (fun (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))))).
21620 (* constant 4335 *)
21621 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))))))).
21622 exact (fun (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))))))).
21625 (* constant 4336 *)
21626 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))).
21627 exact (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))).
21630 (* constant 4337 *)
21631 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)))).
21632 exact (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))).
21635 (* constant 4338 *)
21636 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))))))).
21637 exact (fun (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)))))))).
21640 (* constant 4339 *)
21641 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))))))).
21642 exact (fun (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)))))))).
21645 (* constant 4340 *)
21646 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))).
21647 exact (fun (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))))))).
21650 (* constant 4341 *)
21651 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))))))).
21652 exact (fun (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)))))))).
21655 (* constant 4342 *)
21656 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))))))).
21657 exact (fun (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)))))))).
21660 (* constant 4343 *)
21661 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))).
21662 exact (fun (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))))))).
21665 (* constant 4344 *)
21666 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))))))).
21667 exact (fun (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)))))))).
21670 (* constant 4345 *)
21671 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))))))).
21672 exact (fun (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)))))))).
21675 (* constant 4346 *)
21676 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))).
21677 exact (fun (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))))))).
21680 (* constant 4347 *)
21681 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)))))))).
21682 exact (fun (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)))))))).
21685 (* constant 4348 *)
21686 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)))))))).
21687 exact (fun (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)))))))).
21690 (* constant 4349 *)
21691 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)))).
21692 exact (fun (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))))))).
21695 (* constant 4350 *)
21696 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)))))))).
21697 exact (fun (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)))))))).
21700 (* constant 4351 *)
21701 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))))))).
21702 exact (fun (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)))))))).
21705 (* constant 4352 *)
21706 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))).
21707 exact (fun (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))))))).
21710 (* constant 4353 *)
21711 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)))))))).
21712 exact (fun (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)))))))).
21715 (* constant 4354 *)
21716 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)))))))).
21717 exact (fun (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)))))))).
21720 (* constant 4355 *)
21721 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)))).
21722 exact (fun (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))))))).
21725 (* constant 4356 *)
21726 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)))))))).
21727 exact (fun (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)))))))).
21730 (* constant 4357 *)
21731 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)))))))).
21732 exact (fun (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)))))))).
21735 (* constant 4358 *)
21736 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)))).
21737 exact (fun (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))))))).
21740 (* constant 4359 *)
21741 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)))).
21742 exact (fun (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))).
21745 (* constant 4360 *)
21746 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)))))))).
21747 exact (fun (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)))))))).
21750 (* constant 4361 *)
21751 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)))))))).
21752 exact (fun (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)))))))).
21755 (* constant 4362 *)
21756 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)))).
21757 exact (fun (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))))))).
21760 (* constant 4363 *)
21761 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))).
21762 exact (fun (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)))).
21765 (* constant 4364 *)
21766 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))).
21767 exact (fun (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)))).
21770 (* constant 4365 *)
21771 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))).
21772 exact (fun (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)))).
21775 (* constant 4366 *)
21776 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))).
21777 exact (fun (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))))).
21780 (* constant 4367 *)
21781 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)).
21782 exact (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))).
21785 (* constant 4368 *)
21786 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).
21787 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)).
21790 (* constant 4369 *)
21791 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))))).
21792 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))))))).
21795 (* constant 4370 *)
21796 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)))).
21797 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))))).
21800 (* constant 4371 *)
21801 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))))))).
21802 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))))))).
21805 (* constant 4372 *)
21806 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))))))).
21807 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))))))).
21810 (* constant 4373 *)
21811 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)))))))).
21812 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))))))).
21815 (* constant 4374 *)
21816 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))))).
21817 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))))).
21820 (* constant 4375 *)
21821 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))))))).
21822 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))))))).
21825 (* constant 4376 *)
21826 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))))))))).
21827 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))))))).
21830 (* constant 4377 *)
21831 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))))))))).
21832 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))))))).
21835 (* constant 4378 *)
21836 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))))))))).
21837 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)))))))).
21840 (* constant 4379 *)
21841 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)))))))).
21842 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)))))))).
21845 (* constant 4380 *)
21846 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)))))))).
21847 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)))))))).
21850 (* constant 4381 *)
21851 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)))))).
21852 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))))))).
21855 (* constant 4382 *)
21856 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))))).
21857 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)))))).
21860 (* constant 4383 *)
21861 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))).
21862 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))))).
21865 (* constant 4384 *)
21866 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))))).
21867 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))).
21870 (* constant 4385 *)
21871 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)))))))))))))).
21872 exact (fun (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)))))))))))).
21875 (* constant 4386 *)
21876 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)))))))))))))).
21877 exact (fun (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))))))))))))).
21880 (* constant 4387 *)
21881 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)))))).
21882 exact (fun (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)))))))))))).
21885 (* constant 4388 *)
21886 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))))))))))).
21887 exact (fun (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))))))))).
21890 (* constant 4389 *)
21891 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))))))))))).
21892 exact (fun (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)))))))))).
21895 (* constant 4390 *)
21896 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))))).
21897 exact (fun (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))))))))).
21900 (* constant 4391 *)
21901 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))))).
21902 exact (fun (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))).
21905 (* constant 4392 *)
21906 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)))).
21907 exact (fun (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)))).
21910 (* constant 4393 *)
21911 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)).
21912 exact (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)))).
21915 (* constant 4394 *)
21916 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)).
21917 exact (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)))).
21920 (* constant 4395 *)
21921 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)).
21922 exact (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))).
21925 (* constant 4396 *)
21926 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))).
21927 exact (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))).
21930 (* constant 4397 *)
21931 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)))).
21932 exact (fun (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))))).
21935 (* constant 4398 *)
21936 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))))).
21937 exact (fun (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))))).
21940 (* constant 4399 *)
21941 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)))).
21942 exact (fun (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))))).
21945 (* constant 4400 *)
21946 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))))).
21947 exact (fun (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))))).
21950 (* constant 4401 *)
21951 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)))))).
21952 exact (fun (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))))))).
21955 (* constant 4402 *)
21956 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))).
21957 exact (fun (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)))))).
21960 (* constant 4403 *)
21961 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))).
21962 exact (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))).
21965 (* constant 4404 *)
21966 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)))))))))).
21967 exact (fun (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))))))))))).
21970 (* constant 4405 *)
21971 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)))))))))).
21972 exact (fun (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))))))))))).
21975 (* constant 4406 *)
21976 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)))).
21977 exact (fun (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)))))))))).
21980 (* constant 4407 *)
21981 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)))))))))).
21982 exact (fun (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))))))))))).
21985 (* constant 4408 *)
21986 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)))))))))).
21987 exact (fun (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))))))))))).
21990 (* constant 4409 *)
21991 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)))).
21992 exact (fun (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)))))))))).
21995 (* constant 4410 *)
21996 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)))))))))).
21997 exact (fun (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))))))))))).
22000 (* constant 4411 *)
22001 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)))))))))).
22002 exact (fun (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))))))))))).
22005 (* constant 4412 *)
22006 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)))).
22007 exact (fun (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)))))))))).
22010 (* constant 4413 *)
22011 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))))))))))).
22012 exact (fun (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))))))))))).
22015 (* constant 4414 *)
22016 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))))))))))).
22017 exact (fun (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))))))))))).
22020 (* constant 4415 *)
22021 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))))).
22022 exact (fun (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)))))))))).
22025 (* constant 4416 *)
22026 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))))).
22027 exact (fun (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)))).
22030 (* constant 4417 *)
22031 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))))))))))).
22032 exact (fun (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))))))))))).
22035 (* constant 4418 *)
22036 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))))))))))).
22037 exact (fun (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))))))))))).
22040 (* constant 4419 *)
22041 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))))).
22042 exact (fun (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)))))))))).
22045 (* constant 4420 *)
22046 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)))).
22047 exact (fun (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))))).
22050 (* constant 4421 *)
22051 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)))).
22052 exact (fun (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)))))).
22055 (* constant 4422 *)
22056 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)))).
22057 exact (fun (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))))).
22060 (* constant 4423 *)
22061 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))))).
22062 exact (fun (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))))).
22065 (* constant 4424 *)
22066 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))))).
22067 exact (fun (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)))).
22070 (* constant 4425 *)
22071 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))))).
22072 exact (fun (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))))).
22075 (* constant 4426 *)
22076 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))))))).
22077 exact (fun (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))))))).
22080 (* constant 4427 *)
22081 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))))))).
22082 exact (fun (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))))))).
22085 (* constant 4428 *)
22086 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))))))).
22087 exact (fun (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))))))).
22090 (* constant 4429 *)
22091 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))))))).
22092 exact (fun (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))))))).
22095 (* constant 4430 *)
22096 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))))))).
22097 exact (fun (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))))))).
22100 (* constant 4431 *)
22101 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))))))).
22102 exact (fun (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)))))))).
22105 (* constant 4432 *)
22106 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))))))).
22107 exact (fun (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))))))).
22110 (* constant 4433 *)
22111 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))))))).
22112 exact (fun (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))))))).
22115 (* constant 4434 *)
22116 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))))))).
22117 exact (fun (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)))))))).
22120 (* constant 4435 *)
22121 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))))))).
22122 exact (fun (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)))))))).
22125 (* constant 4436 *)
22126 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))))))))))))))).
22127 exact (fun (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))))))))))))))).
22130 (* constant 4437 *)
22131 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))))))))))))))).
22132 exact (fun (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))))))))))))))).
22135 (* constant 4438 *)
22136 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))))))).
22137 exact (fun (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)))))))))))))).
22140 (* constant 4439 *)
22141 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))))))).
22142 exact (fun (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)))))))).
22145 (* constant 4440 *)
22146 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)).
22147 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))).
22150 (* constant 4441 *)
22151 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))))))).
22152 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))))))).
22155 (* constant 4442 *)
22156 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.
22157 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)))))).
22160 (* constant 4443 *)
22161 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)).
22162 exact (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)).
22165 (* constant 4444 *)
22166 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))))))).
22167 exact (fun (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)))))).
22170 (* constant 4445 *)
22171 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)))))))).
22172 exact (fun (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))))))).
22175 (* constant 4446 *)
22176 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))))).
22177 exact (fun (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)))).
22180 (* constant 4447 *)
22181 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))))).
22182 exact (fun (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)))).
22185 (* constant 4448 *)
22186 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))))))).
22187 exact (fun (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))))))).
22190 (* constant 4449 *)
22191 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)))))))).
22192 exact (fun (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)))))))).
22195 (* constant 4450 *)
22196 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))))))).
22197 exact (fun (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)))))))).
22200 (* constant 4451 *)
22201 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))).
22202 exact (fun (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))))))).
22205 (* constant 4452 *)
22206 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)))))))).
22207 exact (fun (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)))))))).
22210 (* constant 4453 *)
22211 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))))))).
22212 exact (fun (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)))))))).
22215 (* constant 4454 *)
22216 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))).
22217 exact (fun (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))))))).
22220 (* constant 4455 *)
22221 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)))))))).
22222 exact (fun (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)))))))).
22225 (* constant 4456 *)
22226 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)))))))).
22227 exact (fun (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)))))))).
22230 (* constant 4457 *)
22231 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)))).
22232 exact (fun (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))))))).
22235 (* constant 4458 *)
22236 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)))).
22237 exact (fun (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))))).
22240 (* constant 4459 *)
22241 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))).
22242 exact (fun (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))).
22245 (* constant 4460 *)
22246 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))).
22247 exact (fun (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))).
22250 (* constant 4461 *)
22251 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)))))))).
22252 exact (fun (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)))))).
22255 (* constant 4462 *)
22256 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)))))))).
22257 exact (fun (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))))))).
22260 (* constant 4463 *)
22261 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)))).
22262 exact (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)))))).
22265 (* constant 4464 *)
22266 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)))).
22267 exact (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))).
22270 (* constant 4465 *)
22271 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))))))).
22272 exact (fun (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)))))).
22275 (* constant 4466 *)
22276 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))))))).
22277 exact (fun (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))))))).
22280 (* constant 4467 *)
22281 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))).
22282 exact (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)))))).
22285 (* constant 4468 *)
22286 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))).
22287 exact (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)).
22290 (* constant 4469 *)
22291 Definition l_e_st_eq_landau_n_rt_rp_r_1rl : l_e_st_eq_landau_n_rt_rp_r_real.
22292 exact (l_e_st_eq_landau_n_rt_rp_r_realof l_e_st_eq_landau_n_rt_rp_1df).
22295 (* constant 4470 *)
22296 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.
22297 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)).
22300 (* constant 4471 *)
22301 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.
22302 exact (l_e_st_eq_landau_n_rt_rp_r_natrli l_e_st_eq_landau_n_1).
22305 (* constant 4472 *)
22306 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.
22307 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).
22310 (* constant 4473 *)
22311 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))).
22312 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))).
22315 (* constant 4474 *)
22316 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))).
22317 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)))).
22320 (* constant 4475 *)
22321 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).
22322 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))).
22325 (* constant 4476 *)
22326 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)).
22327 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)).
22330 (* constant 4477 *)
22331 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).
22332 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)).
22335 (* constant 4478 *)
22336 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)).
22337 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)).
22340 (* constant 4479 *)
22341 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)))))).
22342 exact (fun (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)))))).
22345 (* constant 4480 *)
22346 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)))))))))).
22347 exact (fun (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))))))))).
22350 (* constant 4481 *)
22351 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)))))))))).
22352 exact (fun (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))))))))).
22355 (* constant 4482 *)
22356 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)))))).
22357 exact (fun (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)))))))).
22360 (* constant 4483 *)
22361 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))))))))))).
22362 exact (fun (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))))))))).
22365 (* constant 4484 *)
22366 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))))))))))).
22367 exact (fun (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))))))))).
22370 (* constant 4485 *)
22371 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))))))).
22372 exact (fun (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)))))))).
22375 (* constant 4486 *)
22376 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))))))).
22377 exact (fun (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))))))).
22380 (* constant 4487 *)
22381 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))))).
22382 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))))).
22385 (* constant 4488 *)
22386 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))))))))))).
22387 exact (fun (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)))))))))).
22390 (* constant 4489 *)
22391 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))))))))))).
22392 exact (fun (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)))))))))))).
22395 (* constant 4490 *)
22396 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))))))))))).
22397 exact (fun (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)))))))))))).
22400 (* constant 4491 *)
22401 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))))))))))).
22402 exact (fun (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)))))))))).
22405 (* constant 4492 *)
22406 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))))))).
22407 exact (fun (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))))))))).
22410 (* constant 4493 *)
22411 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))))))))))).
22412 exact (fun (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)))))))))).
22415 (* constant 4494 *)
22416 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))))))))))).
22417 exact (fun (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)))))))))))).
22420 (* constant 4495 *)
22421 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))))))))))).
22422 exact (fun (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)))))))))))).
22425 (* constant 4496 *)
22426 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))))))))))).
22427 exact (fun (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)))))))))).
22430 (* constant 4497 *)
22431 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))))))).
22432 exact (fun (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))))))))).
22435 (* constant 4498 *)
22436 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))).
22437 exact (fun (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)))).
22440 (* constant 4499 *)
22441 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))).
22442 exact (fun (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)))).
22445 (* constant 4500 *)
22446 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))))).
22447 exact (fun (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)))).
22450 (* constant 4501 *)
22451 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))))).
22452 exact (fun (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)))).
22455 (* constant 4502 *)
22456 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))).
22457 exact (fun (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)))).
22460 (* constant 4503 *)
22461 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))).
22462 exact (fun (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)))).
22465 (* constant 4504 *)
22466 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)))))).
22467 exact (fun (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))))).
22470 (* constant 4505 *)
22471 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))))).
22472 exact (fun (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)))).
22475 (* constant 4506 *)
22476 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)))))))).
22477 exact (fun (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)))))).
22480 (* constant 4507 *)
22481 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)))))))).
22482 exact (fun (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))))))).
22485 (* constant 4508 *)
22486 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)))).
22487 exact (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)))))).
22490 (* constant 4509 *)
22491 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)))).
22492 exact (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)))).
22495 (* constant 4510 *)
22496 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)))).
22497 exact (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))).
22500 (* constant 4511 *)
22501 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))).
22502 exact (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))).
22505 (* constant 4512 *)
22506 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))).
22507 exact (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))).
22510 (* constant 4513 *)
22511 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)))).
22512 exact (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))).
22515 (* constant 4514 *)
22516 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))).
22517 exact (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)))).
22520 (* constant 4515 *)
22521 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)))).
22522 exact (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))).
22525 (* constant 4516 *)
22526 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))))))))).
22527 exact (fun (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))))))))).
22530 (* constant 4517 *)
22531 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))))))))).
22532 exact (fun (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))))))))).
22535 (* constant 4518 *)
22536 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))))).
22537 exact (fun (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)))))))).
22540 (* constant 4519 *)
22541 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))))))))).
22542 exact (fun (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))))))))).
22545 (* constant 4520 *)
22546 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))))))))).
22547 exact (fun (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))))))))).
22550 (* constant 4521 *)
22551 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))))).
22552 exact (fun (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)))))))).
22555 (* constant 4522 *)
22556 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))))).
22557 exact (fun (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))))).
22560 (* constant 4523 *)
22561 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))))).
22562 exact (fun (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)))))).
22565 (* constant 4524 *)
22566 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))).
22567 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))).
22570 (* constant 4525 *)
22571 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))).
22572 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))).
22575 (* constant 4526 *)
22576 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))))))))))).
22577 exact (fun (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))))))))).
22580 (* constant 4527 *)
22581 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))))))))))).
22582 exact (fun (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)))))))))).
22585 (* constant 4528 *)
22586 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))))).
22587 exact (fun (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))))))))).
22590 (* constant 4529 *)
22591 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))))).
22592 exact (fun (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))).
22595 (* constant 4530 *)
22596 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)))).
22597 exact (fun (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)))).
22600 (* constant 4531 *)
22601 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))))))))))).
22602 exact (fun (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))))))))).
22605 (* constant 4532 *)
22606 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))))))))))).
22607 exact (fun (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)))))))))).
22610 (* constant 4533 *)
22611 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))))).
22612 exact (fun (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))))))))).
22615 (* constant 4534 *)
22616 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))))).
22617 exact (fun (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))))).
22620 (* constant 4535 *)
22621 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))))).
22622 exact (fun (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))).
22625 (* constant 4536 *)
22626 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)))).
22627 exact (fun (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)))).
22630 (* constant 4537 *)
22631 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))))).
22632 exact (fun (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)))).
22635 (* constant 4538 *)
22636 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))))).
22637 exact (fun (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))))).
22640 (* constant 4539 *)
22641 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))))).
22642 exact (fun (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))))).
22645 (* constant 4540 *)
22646 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))))).
22647 exact (fun (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))).
22650 (* constant 4541 *)
22651 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)))).
22652 exact (fun (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)))).
22655 (* constant 4542 *)
22656 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))))).
22657 exact (fun (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)))).
22660 (* constant 4543 *)
22661 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))))).
22662 exact (fun (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))).
22665 (* constant 4544 *)
22666 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)))))))))))).
22667 exact (fun (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)))))))))))).
22670 (* constant 4545 *)
22671 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)))))))))))).
22672 exact (fun (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)))))))))))).
22675 (* constant 4546 *)
22676 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)))))).
22677 exact (fun (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))))))))))).
22680 (* constant 4547 *)
22681 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)))))).
22682 exact (fun (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)))))).
22685 (* constant 4548 *)
22686 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)))))))))))).
22687 exact (fun (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)))))))))))).
22690 (* constant 4549 *)
22691 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)))))))))))).
22692 exact (fun (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)))))))))))).
22695 (* constant 4550 *)
22696 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)))))).
22697 exact (fun (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))))))))))).
22700 (* constant 4551 *)
22701 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)))))).
22702 exact (fun (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)))))).
22705 (* constant 4552 *)
22706 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)))))).
22707 exact (fun (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)))))).
22710 (* constant 4553 *)
22711 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)))))).
22712 exact (fun (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)))))).
22715 (* constant 4554 *)
22716 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)))))).
22717 exact (fun (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)))))).
22720 (* constant 4555 *)
22721 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)))))).
22722 exact (fun (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)))))).
22725 (* constant 4556 *)
22726 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)))))).
22727 exact (fun (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)))))).
22730 (* constant 4557 *)
22731 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)))))).
22732 exact (fun (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)))))).
22735 (* constant 4558 *)
22736 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)))))).
22737 exact (fun (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)))))).
22740 (* constant 4559 *)
22741 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)))))).
22742 exact (fun (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)))))).
22745 (* constant 4560 *)
22746 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))))).
22747 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))))).
22750 (* constant 4561 *)
22751 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))))))))))))))).
22752 exact (fun (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)))))))))))))))).
22755 (* constant 4562 *)
22756 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))))))))))))))).
22757 exact (fun (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)))))))))))))))).
22760 (* constant 4563 *)
22761 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))))))).
22762 exact (fun (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))))))))))))))).
22765 (* constant 4564 *)
22766 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)))))))).
22767 exact (fun (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)))))))).
22770 (* constant 4565 *)
22771 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))))))))).
22772 exact (fun (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))))))))).
22775 (* constant 4566 *)
22776 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))))))))).
22777 exact (fun (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))))))))).
22780 (* constant 4567 *)
22781 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)))))))))).
22782 exact (fun (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)))))))))).
22785 (* constant 4568 *)
22786 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)))))))).
22787 exact (fun (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))))))))).
22790 (* constant 4569 *)
22791 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)))).
22792 exact (fun (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))))))).
22795 (* constant 4570 *)
22796 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)))).
22797 exact (fun (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)))).
22800 (* constant 4571 *)
22801 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))).
22802 exact (fun (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)))).
22805 (* constant 4572 *)
22806 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))).
22807 exact (fun (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)))).
22810 (* constant 4573 *)
22811 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))))).
22812 exact (fun (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)))).
22815 (* constant 4574 *)
22816 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))).
22817 exact (fun (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)))).
22820 (* constant 4575 *)
22821 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)))).
22822 exact (fun (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)))).
22825 (* constant 4576 *)
22826 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)))))).
22827 exact (fun (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)))))).
22830 (* constant 4577 *)
22831 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))).
22832 exact (fun (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))).
22835 (* constant 4578 *)
22836 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)))))).
22837 exact (fun (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)))).
22840 (* constant 4579 *)
22841 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)))))))).
22842 exact (fun (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)))))).
22845 (* constant 4580 *)
22846 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))))))).
22847 exact (fun (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)))))).
22850 (* constant 4581 *)
22851 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)))))).
22852 exact (fun (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)))))).
22855 (* constant 4582 *)
22856 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)))))))).
22857 exact (fun (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)))))).
22860 (* constant 4583 *)
22861 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))))))).
22862 exact (fun (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)))))).
22865 (* constant 4584 *)
22866 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)))))).
22867 exact (fun (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)))))).
22870 (* constant 4585 *)
22871 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)))))).
22872 exact (fun (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)))).
22875 (* constant 4586 *)
22876 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)))))))).
22877 exact (fun (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)))))).
22880 (* constant 4587 *)
22881 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))))))).
22882 exact (fun (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)))))).
22885 (* constant 4588 *)
22886 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)))))).
22887 exact (fun (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)))))).
22890 (* constant 4589 *)
22891 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)))))))).
22892 exact (fun (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)))))).
22895 (* constant 4590 *)
22896 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))))))).
22897 exact (fun (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)))))).
22900 (* constant 4591 *)
22901 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)))))).
22902 exact (fun (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)))))).
22905 (* constant 4592 *)
22906 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)))).
22907 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)))).
22910 (* constant 4593 *)
22911 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))).
22912 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)))).
22915 (* constant 4594 *)
22916 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)))).
22917 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))))).
22920 (* constant 4595 *)
22921 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))).
22922 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))))).
22925 (* constant 4596 *)
22926 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))))))).
22927 exact (fun (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))))).
22930 (* constant 4597 *)
22931 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)))))).
22932 exact (fun (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)))))).
22935 (* constant 4598 *)
22936 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))))))).
22937 exact (fun (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))))).
22940 (* constant 4599 *)
22941 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))))).
22942 exact (fun (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)))))).
22945 (* constant 4600 *)
22946 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)))))).
22947 exact (fun (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))))))).
22950 (* constant 4601 *)
22951 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))))).
22952 exact (fun (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))))))).
22955 (* constant 4602 *)
22956 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).
22957 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)).
22960 (* constant 4603 *)
22961 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).
22962 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)).
22965 (* constant 4604 *)
22966 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)))).
22967 exact (fun (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)))).
22970 (* constant 4605 *)
22971 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)))).
22972 exact (fun (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))))).
22975 (* constant 4606 *)
22976 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)))).
22977 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))).
22980 (* constant 4607 *)
22981 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)).
22982 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))).
22985 (* constant 4608 *)
22986 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)).
22987 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)).
22990 (* constant 4609 *)
22991 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).
22992 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))).
22995 (* constant 4610 *)
22996 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)).
22997 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)).
23000 (* constant 4611 *)
23001 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)).
23002 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)).
23005 (* constant 4612 *)
23006 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))))).
23007 exact (fun (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))))))).
23010 (* constant 4613 *)
23011 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))))).
23012 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))))).
23015 (* constant 4614 *)
23016 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)))).
23017 exact (fun (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)))).
23020 (* constant 4615 *)
23021 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)))).
23022 exact (fun (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))).
23025 (* constant 4616 *)
23026 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).
23027 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)).
23030 (* constant 4617 *)
23031 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).
23032 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)).
23035 (* constant 4618 *)
23036 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)))).
23037 exact (fun (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)))).
23040 (* constant 4619 *)
23041 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)))).
23042 exact (fun (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))))).
23045 (* constant 4620 *)
23046 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)))).
23047 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))).
23050 (* constant 4621 *)
23051 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).
23052 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)))).
23055 (* constant 4622 *)
23056 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)).
23057 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)).
23060 (* constant 4623 *)
23061 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)).
23062 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)).
23065 (* constant 4624 *)
23066 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)).
23067 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))).
23070 (* constant 4625 *)
23071 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)).
23072 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)).
23075 (* constant 4626 *)
23076 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))))).
23077 exact (fun (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))))))).
23080 (* constant 4627 *)
23081 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))))).
23082 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))))).
23085 (* constant 4628 *)
23086 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)))).
23087 exact (fun (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))).
23090 (* constant 4629 *)
23091 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)))).
23092 exact (fun (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)))).
23095 (* constant 4630 *)
23096 Definition l_e_st_eq_landau_n_rt_rp_r_2rl : l_e_st_eq_landau_n_rt_rp_r_real.
23097 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).
23100 (* constant 4631 *)
23101 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.
23102 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).
23105 (* constant 4632 *)
23106 Definition l_e_st_eq_landau_n_rt_rp_r_half : l_e_st_eq_landau_n_rt_rp_r_real.
23107 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)).
23110 (* constant 4633 *)
23111 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.
23112 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).
23115 (* constant 4634 *)
23116 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)).
23117 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)).
23120 (* constant 4635 *)
23121 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).
23122 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)).
23125 (* constant 4636 *)
23126 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))))).
23127 exact (fun (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))).
23130 (* constant 4637 *)
23131 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))))).
23132 exact (fun (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)))).
23135 (* constant 4638 *)
23136 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))))).
23137 exact (fun (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))).
23140 (* constant 4639 *)
23141 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))).
23142 exact (fun (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)))).
23145 (* constant 4640 *)
23146 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)))).
23147 exact (fun (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)))))).
23150 (* constant 4641 *)
23151 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))).
23152 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))))).
23155 (* constant 4642 *)
23156 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))).
23157 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))))).
23160 (* constant 4643 *)
23161 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))).
23162 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)))).
23165 (* constant 4644 *)
23166 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))))))))))).
23167 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)))))))))))).
23170 (* constant 4645 *)
23171 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))))))))))).
23172 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)))))))))))).
23175 (* constant 4646 *)
23176 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))))))))))).
23177 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))))))))))).
23180 (* constant 4647 *)
23181 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))))))))))).
23182 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)))))))))))).
23185 (* constant 4648 *)
23186 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))))))))))).
23187 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)))))))))))).
23190 (* constant 4649 *)
23191 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)))))))))))).
23192 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)))))))))))).
23195 (* constant 4650 *)
23196 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))))))))))).
23197 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))))))))))))).
23200 (* constant 4651 *)
23201 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))))))))))).
23202 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))))))))))).
23205 (* constant 4652 *)
23206 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))))))))))).
23207 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)))))))))))).
23210 (* constant 4653 *)
23211 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)))))))))).
23212 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))))))))))).
23215 (* constant 4654 *)
23216 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))))))).
23217 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)))))))))).
23220 (* constant 4655 *)
23221 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))))))))).
23222 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))))))))).
23225 (* constant 4656 *)
23226 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))))))))).
23227 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))))))))).
23230 (* constant 4657 *)
23231 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))))))))).
23232 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)))))))))).
23235 (* constant 4658 *)
23236 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))))))))).
23237 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)))))))))).
23240 (* constant 4659 *)
23241 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)))))))))))).
23242 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))))))))))).
23245 (* constant 4660 *)
23246 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))))))))))).
23247 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))))))))))).
23250 (* constant 4661 *)
23251 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)))))))))))).
23252 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))))))))))).
23255 (* constant 4662 *)
23256 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))))))))))).
23257 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))))))))))).
23260 (* constant 4663 *)
23261 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)))))))))))).
23262 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))))))))))).
23265 (* constant 4664 *)
23266 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))))))))).
23267 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)))))))))).
23270 (* constant 4665 *)
23271 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))))))))).
23272 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))))))))))).
23275 (* constant 4666 *)
23276 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)))))))))).
23277 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))))))))))).
23280 (* constant 4667 *)
23281 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))))))))))).
23282 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))))))))))).
23285 (* constant 4668 *)
23286 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))))))))))).
23287 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)))))))))))).
23290 (* constant 4669 *)
23291 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))))))))))).
23292 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)))))))))))).
23295 (* constant 4670 *)
23296 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))))))))))).
23297 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))))))))))))).
23300 (* constant 4671 *)
23301 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)))))))))))).
23302 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))))))))))))).
23305 (* constant 4672 *)
23306 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)))))))))).
23307 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))))))))))).
23310 (* constant 4673 *)
23311 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)))))))))))))).
23312 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)))))))))))))).
23315 (* constant 4674 *)
23316 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))))))))))))).
23317 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)))))))))))))).
23320 (* constant 4675 *)
23321 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))))))))).
23322 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))))))))))))).
23325 (* constant 4676 *)
23326 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)))))))))))).
23327 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))))))))))))).
23330 (* constant 4677 *)
23331 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)))))))))))).
23332 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))))))))))))).
23335 (* constant 4678 *)
23336 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))))))))).
23337 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)))))))))).
23340 (* constant 4679 *)
23341 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)))))))))).
23342 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)))))))))).
23345 (* constant 4680 *)
23346 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)))))))))))).
23347 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)))))))))))).
23350 (* constant 4681 *)
23351 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))))))))))))).
23352 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))))))))))))).
23355 (* constant 4682 *)
23356 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))))))))))))).
23357 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))))))))))))).
23360 (* constant 4683 *)
23361 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)))))))))))).
23362 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))))))))))))).
23365 (* constant 4684 *)
23366 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))))))))))))).
23367 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))))))))))))).
23370 (* constant 4685 *)
23371 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))))))))))))).
23372 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))))))))))))))).
23375 (* constant 4686 *)
23376 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)))))))))))).
23377 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))))))))))))).
23380 (* constant 4687 *)
23381 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))))))))))).
23382 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)))))))))))).
23385 (* constant 4688 *)
23386 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))))))))))).
23387 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)))))))))))).
23390 (* constant 4689 *)
23391 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))))))))))).
23392 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)))))))))))).
23395 (* constant 4690 *)
23396 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)))))))))))).
23397 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)))))))))))).
23400 (* constant 4691 *)
23401 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)))))))))))).
23402 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)))))))))))).
23405 (* constant 4692 *)
23406 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))))))))))).
23407 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))))))))))))).
23410 (* constant 4693 *)
23411 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)))))))))).
23412 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))))))))))).
23415 (* constant 4694 *)
23416 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)))))))))).
23417 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)))))))))).
23420 (* constant 4695 *)
23421 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)))))))).
23422 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))))))))).
23425 (* constant 4696 *)
23426 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))))))).
23427 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)))))))).
23430 (* constant 4697 *)
23431 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))))))).
23432 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)))))))).
23435 (* constant 4698 *)
23436 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)))))))))).
23437 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))))))))).
23440 (* constant 4699 *)
23441 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))))))))).
23442 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))))))))).
23445 (* constant 4700 *)
23446 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)))))))))).
23447 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))))))))).
23450 (* constant 4701 *)
23451 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))))))))).
23452 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))))))))).
23455 (* constant 4702 *)
23456 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)))))))))).
23457 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)))))))))).
23460 (* constant 4703 *)
23461 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)))))))))).
23462 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))))))))))).
23465 (* constant 4704 *)
23466 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)))))))))).
23467 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)))))))))).
23470 (* constant 4705 *)
23471 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)))))))).
23472 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))))))))).
23475 (* constant 4706 *)
23476 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)))))))))).
23477 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))))))))))).
23480 (* constant 4707 *)
23481 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)))))))))).
23482 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)))))))))).
23485 (* constant 4708 *)
23486 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)))))))).
23487 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))))))))).
23490 (* constant 4709 *)
23491 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)))))))))))).
23492 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)))))))))))).
23495 (* constant 4710 *)
23496 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))))))))))).
23497 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))))))))))))).
23500 (* constant 4711 *)
23501 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)))))))))).
23502 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)))))))))).
23505 (* constant 4712 *)
23506 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))))))))).
23507 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)))))))))).
23510 (* constant 4713 *)
23511 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))))))))))).
23512 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))))))))))).
23515 (* constant 4714 *)
23516 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)))))))))))).
23517 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)))))))))).
23520 (* constant 4715 *)
23521 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)))))))))).
23522 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))))))))).
23525 (* constant 4716 *)
23526 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)))))))).
23527 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)))))))).
23530 (* constant 4717 *)
23531 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))))))))))).
23532 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)))))))))))).
23535 (* constant 4718 *)
23536 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)))))))))))).
23537 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)))))))))))).
23540 (* constant 4719 *)
23541 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))))))))))).
23542 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)))))))))))).
23545 (* constant 4720 *)
23546 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))))))))))).
23547 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)))))))))))).
23550 (* constant 4721 *)
23551 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)))))))))))).
23552 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)))))))))))).
23555 (* constant 4722 *)
23556 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))))))))))).
23557 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)))))))))))).
23560 (* constant 4723 *)
23561 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)))))))))).
23562 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))))))))))).
23565 (* constant 4724 *)
23566 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)))))))))).
23567 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)))))))))).
23570 (* constant 4725 *)
23571 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)))))))).
23572 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))))))))).
23575 (* constant 4726 *)
23576 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)))))))))))).
23577 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)))))))))).
23580 (* constant 4727 *)
23581 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))))))))))).
23582 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))))))))))).
23585 (* constant 4728 *)
23586 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)))))))))).
23587 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))))))))))).
23590 (* constant 4729 *)
23591 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)))))))))))).
23592 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)))))))))).
23595 (* constant 4730 *)
23596 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))))))))))).
23597 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))))))))))).
23600 (* constant 4731 *)
23601 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)))))))))).
23602 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))))))))))).
23605 (* constant 4732 *)
23606 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)))))))).
23607 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)))))))))).
23610 (* constant 4733 *)
23611 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))))))))).
23612 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))))))))).
23615 (* constant 4734 *)
23616 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)))))))).
23617 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)))))))).
23620 (* constant 4735 *)
23621 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))))))).
23622 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))))))).
23625 (* constant 4736 *)
23626 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))))))).
23627 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))))))).
23630 (* constant 4737 *)
23631 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)))))))))).
23632 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)))))).
23635 (* constant 4738 *)
23636 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)))))))))).
23637 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)))))).
23640 (* constant 4739 *)
23641 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)))))).
23642 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))))))).
23645 (* constant 4740 *)
23646 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)))))))).
23647 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)))))))).
23650 (* constant 4741 *)
23651 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)))))))).
23652 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)))))))).
23655 (* constant 4742 *)
23656 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)).
23657 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)).
23660 (* constant 4743 *)
23661 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)).
23662 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)).
23665 (* constant 4744 *)
23666 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)))).
23667 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)).
23670 (* constant 4745 *)
23671 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)))).
23672 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))).
23675 (* constant 4746 *)
23676 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)))).
23677 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)).
23680 (* constant 4747 *)
23681 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)))).
23682 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))).
23685 (* constant 4748 *)
23686 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)))).
23687 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))).
23690 (* constant 4749 *)
23691 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))).
23692 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)))).
23695 (* constant 4750 *)
23696 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)))).
23697 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)))))).
23700 (* constant 4751 *)
23701 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)))).
23702 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))).
23705 (* constant 4752 *)
23706 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))))).
23707 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)))).
23710 (* constant 4753 *)
23711 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))).
23712 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))).
23715 (* constant 4754 *)
23716 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)))).
23717 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)))).
23720 (* constant 4755 *)
23721 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)))).
23722 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)))).
23725 (* constant 4756 *)
23726 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)))).
23727 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))).
23730 (* constant 4757 *)
23731 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)))).
23732 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)))).
23735 (* constant 4758 *)
23736 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)))).
23737 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)))).
23740 (* constant 4759 *)
23741 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))).
23742 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)))).
23745 (* constant 4760 *)
23746 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)))).
23747 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))).
23750 (* constant 4761 *)
23751 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)))).
23752 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)))).
23755 (* constant 4762 *)
23756 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)))).
23757 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)))).
23760 (* constant 4763 *)
23761 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)))).
23762 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)))).
23765 (* constant 4764 *)
23766 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))))).
23767 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))))).
23770 (* constant 4765 *)
23771 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))))).
23772 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))))).
23775 (* constant 4766 *)
23776 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))).
23777 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)))).
23780 (* constant 4767 *)
23781 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))))).
23782 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))))).
23785 (* constant 4768 *)
23786 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))))).
23787 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))))).
23790 (* constant 4769 *)
23791 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))).
23792 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)))).
23795 (* constant 4770 *)
23796 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))))))))).
23797 exact (fun (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))))))))).
23800 (* constant 4771 *)
23801 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))))))))).
23802 exact (fun (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))))))))).
23805 (* constant 4772 *)
23806 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))))).
23807 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)))))))).
23810 (* constant 4773 *)
23811 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))))).
23812 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))))).
23815 (* constant 4774 *)
23816 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))))))))).
23817 exact (fun (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))))))))).
23820 (* constant 4775 *)
23821 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))))))))).
23822 exact (fun (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))))))))).
23825 (* constant 4776 *)
23826 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))))).
23827 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)))))))).
23830 (* constant 4777 *)
23831 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))).
23832 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))).
23835 (* constant 4778 *)
23836 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))).
23837 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)))).
23840 (* constant 4779 *)
23841 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)).
23842 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))).
23845 (* constant 4780 *)
23846 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)))))).
23847 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)))))).
23850 (* constant 4781 *)
23851 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)))))).
23852 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))))).
23855 (* constant 4782 *)
23856 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)))))).
23857 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)))))).
23860 (* constant 4783 *)
23861 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)))))).
23862 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)))))).
23865 (* constant 4784 *)
23866 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)))))).
23867 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)))))).
23870 (* constant 4785 *)
23871 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))))).
23872 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)))))).
23875 (* constant 4786 *)
23876 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))))).
23877 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))).
23880 (* constant 4787 *)
23881 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))))).
23882 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))).
23885 (* constant 4788 *)
23886 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)))).
23887 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))).
23890 (* constant 4789 *)
23891 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)))).
23892 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))).
23895 (* constant 4790 *)
23896 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))))).
23897 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))).
23900 (* constant 4791 *)
23901 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))))).
23902 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))).
23905 (* constant 4792 *)
23906 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)))).
23907 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))).
23910 (* constant 4793 *)
23911 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)))).
23912 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))).
23915 (* constant 4794 *)
23916 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)))))))))))).
23917 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)))))))))))).
23920 (* constant 4795 *)
23921 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)))))))))))).
23922 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)))))))))))).
23925 (* constant 4796 *)
23926 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))))))))))).
23927 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)))))))))))).
23930 (* constant 4797 *)
23931 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))))))))))).
23932 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)))))))))))).
23935 (* constant 4798 *)
23936 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))))))))))).
23937 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))))))))))))).
23940 (* constant 4799 *)
23941 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))))).
23942 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))))))))))).
23945 (* constant 4800 *)
23946 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))))).
23947 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))))).
23950 (* constant 4801 *)
23951 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)))))).
23952 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)))))).
23955 (* constant 4802 *)
23956 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)))))))).
23957 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))))))).
23960 (* constant 4803 *)
23961 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)))))).
23962 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))))))).
23965 (* constant 4804 *)
23966 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))))))).
23967 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))))))).
23970 (* constant 4805 *)
23971 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)))))))).
23972 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))))))).
23975 (* constant 4806 *)
23976 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)))))).
23977 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)))))).
23980 (* constant 4807 *)
23981 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)))).
23982 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)))).
23985 (* constant 4808 *)
23986 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)))).
23987 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))).
23990 (* constant 4809 *)
23991 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)).
23992 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))).
23995 (* constant 4810 *)
23996 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))).
23997 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))).
24000 (* constant 4811 *)
24001 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)))).
24002 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))).
24005 (* constant 4812 *)
24006 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)).
24007 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))).
24010 (* constant 4813 *)
24011 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)))))).
24012 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)))))).
24015 (* constant 4814 *)
24016 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))))).
24017 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))))))).
24020 (* constant 4815 *)
24021 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)))))).
24022 exact (fun (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)))))).
24025 (* constant 4816 *)
24026 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))).
24027 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)))))).
24030 (* constant 4817 *)
24031 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))).
24032 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))))).
24035 (* constant 4818 *)
24036 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)))).
24037 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)))).
24040 (* constant 4819 *)
24041 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))))).
24042 exact (fun (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)))))).
24045 (* constant 4820 *)
24046 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))))).
24047 exact (fun (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))))).
24050 (* constant 4821 *)
24051 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))))).
24052 exact (fun (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)))))).
24055 (* constant 4822 *)
24056 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))))).
24057 exact (fun (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))))).
24060 (* constant 4823 *)
24061 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)).
24062 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))).
24065 (* constant 4824 *)
24066 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)).
24067 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)).
24070 (* constant 4825 *)
24071 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))).
24072 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))))).
24075 (* constant 4826 *)
24076 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))).
24077 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))))).
24080 (* constant 4827 *)
24081 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)).
24082 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)))).
24085 (* constant 4828 *)
24086 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).
24087 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))).
24090 (* constant 4829 *)
24091 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))).
24092 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)).
24095 (* constant 4830 *)
24096 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)))))).
24097 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))))))).
24100 (* constant 4831 *)
24101 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))))).
24102 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)))))).
24105 (* constant 4832 *)
24106 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))))).
24107 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)))))).
24110 (* constant 4833 *)
24111 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)))))).
24112 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)))))).
24115 (* constant 4834 *)
24116 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)))))).
24117 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))))).
24120 (* constant 4835 *)
24121 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)))))).
24122 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)))))).
24125 (* constant 4836 *)
24126 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))))).
24127 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)))))).
24130 (* constant 4837 *)
24131 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)))))).
24132 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)))))).
24135 (* constant 4838 *)
24136 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))))))).
24137 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)))))).
24140 (* constant 4839 *)
24141 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)))))).
24142 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))))))).
24145 (* constant 4840 *)
24146 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))))))).
24147 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)))))))).
24150 (* constant 4841 *)
24151 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)))))).
24152 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)))))).
24155 (* constant 4842 *)
24156 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))))))).
24157 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)))))).
24160 (* constant 4843 *)
24161 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))))))).
24162 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))))))).
24165 (* constant 4844 *)
24166 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))))))))).
24167 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))))))))).
24170 (* constant 4845 *)
24171 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))))))))).
24172 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))))))))).
24175 (* constant 4846 *)
24176 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))))))))).
24177 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))))))))).
24180 (* constant 4847 *)
24181 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)))))))).
24182 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))))))))).
24185 (* constant 4848 *)
24186 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)))))))).
24187 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))))))).
24190 (* constant 4849 *)
24191 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)))))))).
24192 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)))))))).
24195 (* constant 4850 *)
24196 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)))))))).
24197 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)))))))).
24200 (* constant 4851 *)
24201 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)))))))).
24202 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)))))))).
24205 (* constant 4852 *)
24206 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))))))))).
24207 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)))))))).
24210 (* constant 4853 *)
24211 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)))))))).
24212 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)))))))).
24215 (* constant 4854 *)
24216 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)))))))).
24217 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)))))))).
24220 (* constant 4855 *)
24221 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))))))).
24222 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))))))).
24225 (* constant 4856 *)
24226 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)))))).
24227 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))))))).
24230 (* constant 4857 *)
24231 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)))))))).
24232 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))))))).
24235 (* constant 4858 *)
24236 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)))))))).
24237 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)))))))).
24240 (* constant 4859 *)
24241 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)))))))).
24242 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)))))))).
24245 (* constant 4860 *)
24246 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)))))))).
24247 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)))))))).
24250 (* constant 4861 *)
24251 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)))))))).
24252 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))))))).
24255 (* constant 4862 *)
24256 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))))))).
24257 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))))))).
24260 (* constant 4863 *)
24261 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))))))).
24262 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))))))).
24265 (* constant 4864 *)
24266 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))))))).
24267 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))))))).
24270 (* constant 4865 *)
24271 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))))))).
24272 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))))))).
24275 (* constant 4866 *)
24276 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))))))))).
24277 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))))))))).
24280 (* constant 4867 *)
24281 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))))))))).
24282 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))))))))).
24285 (* constant 4868 *)
24286 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)))))))).
24287 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)))))))).
24290 (* constant 4869 *)
24291 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))))))).
24292 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)))))))).
24295 (* constant 4870 *)
24296 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))))))))).
24297 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)))))))).
24300 (* constant 4871 *)
24301 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))))))))).
24302 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)))))))).
24305 (* constant 4872 *)
24306 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)))))))).
24307 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)))))))).
24310 (* constant 4873 *)
24311 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)))))))).
24312 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)))))))).
24315 (* constant 4874 *)
24316 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))))))))).
24317 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)))))))).
24320 (* constant 4875 *)
24321 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))))))))).
24322 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))))))))).
24325 (* constant 4876 *)
24326 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))))))).
24327 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)))))))).
24330 (* constant 4877 *)
24331 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))))))))).
24332 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)))))))))).
24335 (* constant 4878 *)
24336 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))))))).
24337 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)))))))).
24340 (* constant 4879 *)
24341 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)))))).
24342 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)))))))))).
24345 (* constant 4880 *)
24346 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))))))).
24347 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))))))))))))))))).
24350 (* constant 4881 *)
24351 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)))))))).
24352 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))))))))).
24355 (* constant 4882 *)
24356 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)))))).
24357 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))))))))))).
24360 (* constant 4883 *)
24361 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)))))).
24362 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))))))))))))))).
24365 (* constant 4884 *)
24366 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))))))))).
24367 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))))))))))).
24370 (* constant 4885 *)
24371 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)))))))).
24372 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))))))))).
24375 (* constant 4886 *)
24376 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))))))).
24377 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)))))))).
24380 (* constant 4887 *)
24381 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)))))).
24382 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)))))))))).
24385 (* constant 4888 *)
24386 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)))))).
24387 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))))))).
24390 (* constant 4889 *)
24391 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)))))))).
24392 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))))))))).
24395 (* constant 4890 *)
24396 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))))))))).
24397 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))))))))).
24400 (* constant 4891 *)
24401 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))))))))).
24402 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))))))))).
24405 (* constant 4892 *)
24406 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))))))))))))).
24407 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))))))))))).
24410 (* constant 4893 *)
24411 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)))))))))))).
24412 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)))))))))))).
24415 (* constant 4894 *)
24416 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)))))))))))).
24417 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))))))))))))).
24420 (* constant 4895 *)
24421 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)))))))))))).
24422 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)))))))))))).
24425 (* constant 4896 *)
24426 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)))))))))))).
24427 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))))))))))))).
24430 (* constant 4897 *)
24431 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)))))))))))).
24432 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)))))))))))).
24435 (* constant 4898 *)
24436 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))))))))))).
24437 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)))))))))))).
24440 (* constant 4899 *)
24441 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))))))))).
24442 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))))))))))).
24445 (* constant 4900 *)
24446 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))))))))))).
24447 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))))))))))).
24450 (* constant 4901 *)
24451 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))))))))))))).
24452 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)))))))))))).
24455 (* constant 4902 *)
24456 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))))))))))))))).
24457 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)))))))))))).
24460 (* constant 4903 *)
24461 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))))))))))))).
24462 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))))))))))))).
24465 (* constant 4904 *)
24466 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)))))))))))))).
24467 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)))))))))))))).
24470 (* constant 4905 *)
24471 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))))))))))))))).
24472 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)))))))))))))).
24475 (* constant 4906 *)
24476 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))))))))))))))).
24477 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))))))))))))).
24480 (* constant 4907 *)
24481 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)))))))))))))).
24482 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))))))))))))).
24485 (* constant 4908 *)
24486 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)))))))))))).
24487 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))))))))))))).
24490 (* constant 4909 *)
24491 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)))))))))).
24492 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)))))))))))).
24495 (* constant 4910 *)
24496 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)))))))))).
24497 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)))))))))).
24500 (* constant 4911 *)
24501 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)))))))))).
24502 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))))))))))).
24505 (* constant 4912 *)
24506 Definition l_e_st_eq_landau_n_rt_rp_r_c_complex : Type.
24507 exact (l_e_st_eq_landau_n_pair1type l_e_st_eq_landau_n_rt_rp_r_real).
24510 (* constant 4913 *)
24511 Definition l_e_st_eq_landau_n_rt_rp_r_c_cx : Type.
24512 exact l_e_st_eq_landau_n_rt_rp_r_c_complex.
24515 (* constant 4914 *)
24516 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)).
24517 exact (fun (x: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)).
24520 (* constant 4915 *)
24521 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)).
24522 exact (fun (x: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))).
24525 (* constant 4916 *)
24526 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).
24527 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).
24530 (* constant 4917 *)
24531 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).
24532 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).
24535 (* constant 4918 *)
24536 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).
24537 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).
24540 (* constant 4919 *)
24541 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)).
24542 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)).
24545 (* constant 4920 *)
24546 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).
24547 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).
24550 (* constant 4921 *)
24551 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).
24552 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).
24555 (* constant 4922 *)
24556 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)).
24557 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)).
24560 (* constant 4923 *)
24561 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)))).
24562 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)).
24565 (* constant 4924 *)
24566 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)).
24567 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)).
24570 (* constant 4925 *)
24571 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)))).
24572 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)).
24575 (* constant 4926 *)
24576 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).
24577 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).
24580 (* constant 4927 *)
24581 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))).
24582 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).
24585 (* constant 4928 *)
24586 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)))).
24587 exact (fun (x:l_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))).
24590 (* constant 4929 *)
24591 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)))).
24592 exact (fun (x:l_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))).
24595 (* constant 4930 *)
24596 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))))).
24597 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)))).
24600 (* constant 4931 *)
24601 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))))).
24602 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)))).
24605 (* constant 4932 *)
24606 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))))))).
24607 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))))))).
24610 (* constant 4933 *)
24611 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).
24612 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).
24615 (* constant 4934 *)
24616 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))).
24617 exact (fun (x:l_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))).
24620 (* constant 4935 *)
24621 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))))).
24622 exact (fun (x:l_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))))).
24625 (* constant 4936 *)
24626 Definition l_e_st_eq_landau_n_rt_rp_r_c_0c : l_e_st_eq_landau_n_rt_rp_r_c_complex.
24627 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).
24630 (* constant 4937 *)
24631 Definition l_e_st_eq_landau_n_rt_rp_r_c_1c : l_e_st_eq_landau_n_rt_rp_r_c_complex.
24632 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).
24635 (* constant 4938 *)
24636 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)).
24637 exact (fun (x: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)))).
24640 (* constant 4939 *)
24641 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)))))).
24642 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)))))).
24645 (* constant 4940 *)
24646 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)))))).
24647 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))))).
24650 (* constant 4941 *)
24651 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)))))).
24652 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))))).
24655 (* constant 4942 *)
24656 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)))).
24657 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)))).
24660 (* constant 4943 *)
24661 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))))).
24662 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))))).
24665 (* constant 4944 *)
24666 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))))).
24667 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)))).
24670 (* constant 4945 *)
24671 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))))).
24672 exact (fun (x:l_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)))).
24675 (* constant 4946 *)
24676 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))))).
24677 exact (fun (x:l_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)))).
24680 (* constant 4947 *)
24681 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))))))).
24682 exact (fun (x:l_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))))))).
24685 (* constant 4948 *)
24686 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))).
24687 exact (fun (x: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)))).
24690 (* constant 4949 *)
24691 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))).
24692 exact (fun (x: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)).
24695 (* constant 4950 *)
24696 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).
24697 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)).
24700 (* constant 4951 *)
24701 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)).
24702 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)).
24705 (* constant 4952 *)
24706 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).
24707 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)).
24710 (* constant 4953 *)
24711 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)).
24712 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)).
24715 (* constant 4954 *)
24716 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))))).
24717 exact (fun (x:l_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)))))).
24720 (* constant 4955 *)
24721 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))))).
24722 exact (fun (x:l_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))).
24725 (* constant 4956 *)
24726 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)))).
24727 exact (fun (x:l_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)))).
24730 (* constant 4957 *)
24731 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))))).
24732 exact (fun (x:l_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))))).
24735 (* constant 4958 *)
24736 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))))).
24737 exact (fun (x:l_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))))).
24740 (* constant 4959 *)
24741 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)))))).
24742 exact (fun (x:l_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))))).
24745 (* constant 4960 *)
24746 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)))))).
24747 exact (fun (x:l_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))))).
24750 (* constant 4961 *)
24751 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))))))).
24752 exact (fun (x:l_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)))))).
24755 (* constant 4962 *)
24756 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)).
24757 exact (fun (x: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))).
24760 (* constant 4963 *)
24761 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))).
24762 exact (fun (x: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))).
24765 (* constant 4964 *)
24766 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))).
24767 exact (fun (x: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))).
24770 (* constant 4965 *)
24771 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)).
24772 exact (fun (x: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)))).
24775 (* constant 4966 *)
24776 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)))))).
24777 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)))))).
24780 (* constant 4967 *)
24781 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)))))).
24782 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))))).
24785 (* constant 4968 *)
24786 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)))))).
24787 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))))).
24790 (* constant 4969 *)
24791 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)))).
24792 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)))).
24795 (* constant 4970 *)
24796 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))))).
24797 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))))).
24800 (* constant 4971 *)
24801 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))))).
24802 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)))).
24805 (* constant 4972 *)
24806 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))))).
24807 exact (fun (x:l_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)))).
24810 (* constant 4973 *)
24811 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))))).
24812 exact (fun (x:l_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)))).
24815 (* constant 4974 *)
24816 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))))))).
24817 exact (fun (x:l_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))))))).
24820 (* constant 4975 *)
24821 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))))).
24822 exact (fun (x:l_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)))).
24825 (* constant 4976 *)
24826 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)))).
24827 exact (fun (x:l_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))))).
24830 (* constant 4977 *)
24831 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))))).
24832 exact (fun (x:l_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))))).
24835 (* constant 4978 *)
24836 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)))).
24837 exact (fun (x:l_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))))).
24840 (* constant 4979 *)
24841 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)).
24842 exact (fun (x: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)).
24845 (* constant 4980 *)
24846 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))).
24847 exact (fun (x:l_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)))).
24850 (* constant 4981 *)
24851 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))).
24852 exact (fun (x:l_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)))).
24855 (* constant 4982 *)
24856 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)))).
24857 exact (fun (x:l_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)))).
24860 (* constant 4983 *)
24861 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)))).
24862 exact (fun (x:l_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)))).
24865 (* constant 4984 *)
24866 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))).
24867 exact (fun (x:l_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)))).
24870 (* constant 4985 *)
24871 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))).
24872 exact (fun (x:l_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)))).
24875 (* constant 4986 *)
24876 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))).
24877 exact (fun (x:l_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)))).
24880 (* constant 4987 *)
24881 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))).
24882 exact (fun (x:l_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)))).
24885 (* constant 4988 *)
24886 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).
24887 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).
24890 (* constant 4989 *)
24891 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)))).
24892 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))).
24895 (* constant 4990 *)
24896 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)).
24897 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)).
24900 (* constant 4991 *)
24901 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)))).
24902 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))))).
24905 (* constant 4992 *)
24906 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)))).
24907 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))).
24910 (* constant 4993 *)
24911 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)))).
24912 exact (fun (x:l_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))).
24915 (* constant 4994 *)
24916 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).
24917 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)).
24920 (* constant 4995 *)
24921 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))).
24922 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)).
24925 (* constant 4996 *)
24926 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))).
24927 exact (fun (x:l_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)))).
24930 (* constant 4997 *)
24931 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)))).
24932 exact (fun (x:l_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)))).
24935 (* constant 4998 *)
24936 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)))).
24937 exact (fun (x:l_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)))).
24940 (* constant 4999 *)
24941 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))).
24942 exact (fun (x:l_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)))).
24945 (* constant 5000 *)
24946 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).
24947 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)))).
24950 (* constant 5001 *)
24951 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).
24952 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).
24955 (* constant 5002 *)
24956 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).
24957 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)).
24960 (* constant 5003 *)
24961 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)))).
24962 exact (fun (x: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)))).
24965 (* constant 5004 *)
24966 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)))).
24967 exact (fun (x: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))).
24970 (* constant 5005 *)
24971 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)))).
24972 exact (fun (x: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)))).
24975 (* constant 5006 *)
24976 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))).
24977 exact (fun (x: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))).
24980 (* constant 5007 *)
24981 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))).
24982 exact (fun (x: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)))).
24985 (* constant 5008 *)
24986 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))).
24987 exact (fun (x: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))).
24990 (* constant 5009 *)
24991 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)))).
24992 exact (fun (x: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))).
24995 (* constant 5010 *)
24996 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)).
24997 exact (fun (x: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))))).
25000 (* constant 5011 *)
25001 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)).
25002 exact (fun (x: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)))).
25005 (* constant 5012 *)
25006 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)).
25007 exact (fun (x: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)))).
25010 (* constant 5013 *)
25011 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)))))).
25012 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)))))).
25015 (* constant 5014 *)
25016 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)))))).
25017 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)))))).
25020 (* constant 5015 *)
25021 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))))))).
25022 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))))).
25025 (* constant 5016 *)
25026 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)))))).
25027 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))))).
25030 (* constant 5017 *)
25031 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)))))).
25032 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))))).
25035 (* constant 5018 *)
25036 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)))))).
25037 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))))).
25040 (* constant 5019 *)
25041 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))))))).
25042 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)))).
25045 (* constant 5020 *)
25046 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)))).
25047 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)))).
25050 (* constant 5021 *)
25051 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))))).
25052 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))))).
25055 (* constant 5022 *)
25056 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))))).
25057 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))))).
25060 (* constant 5023 *)
25061 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)))))).
25062 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)))).
25065 (* constant 5024 *)
25066 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))))).
25067 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)))).
25070 (* constant 5025 *)
25071 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))))).
25072 exact (fun (x:l_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)))).
25075 (* constant 5026 *)
25076 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))))).
25077 exact (fun (x:l_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)))).
25080 (* constant 5027 *)
25081 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))))))).
25082 exact (fun (x:l_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))))))).
25085 (* constant 5028 *)
25086 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))).
25087 exact (fun (x: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)))).
25090 (* constant 5029 *)
25091 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))).
25092 exact (fun (x: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))))).
25095 (* constant 5030 *)
25096 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))).
25097 exact (fun (x: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))).
25100 (* constant 5031 *)
25101 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))).
25102 exact (fun (x: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)).
25105 (* constant 5032 *)
25106 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)).
25107 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))).
25110 (* constant 5033 *)
25111 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)).
25112 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))).
25115 (* constant 5034 *)
25116 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).
25117 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))).
25120 (* constant 5035 *)
25121 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)).
25122 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)))).
25125 (* constant 5036 *)
25126 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).
25127 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)).
25130 (* constant 5037 *)
25131 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).
25132 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)).
25135 (* constant 5038 *)
25136 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))).
25137 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))))).
25140 (* constant 5039 *)
25141 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)))).
25142 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))))).
25145 (* constant 5040 *)
25146 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))))).
25147 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))))).
25150 (* constant 5041 *)
25151 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))))).
25152 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))))).
25155 (* constant 5042 *)
25156 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)))).
25157 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)))).
25160 (* constant 5043 *)
25161 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))).
25162 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))).
25165 (* constant 5044 *)
25166 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))).
25167 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))).
25170 (* constant 5045 *)
25171 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))).
25172 exact (fun (x:l_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))).
25175 (* constant 5046 *)
25176 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))).
25177 exact (fun (x:l_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))).
25180 (* constant 5047 *)
25181 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))).
25182 exact (fun (x:l_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)))))).
25185 (* constant 5048 *)
25186 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))).
25187 exact (fun (x:l_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))))).
25190 (* constant 5049 *)
25191 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))).
25192 exact (fun (x:l_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)))).
25195 (* constant 5050 *)
25196 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))).
25197 exact (fun (x:l_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)))).
25200 (* constant 5051 *)
25201 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))))).
25202 exact (fun (x:l_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)))).
25205 (* constant 5052 *)
25206 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)))).
25207 exact (fun (x:l_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))))).
25210 (* constant 5053 *)
25211 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)))).
25212 exact (fun (x:l_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))))).
25215 (* constant 5054 *)
25216 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)))).
25217 exact (fun (x:l_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)))))).
25220 (* constant 5055 *)
25221 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)).
25222 exact (fun (x: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))).
25225 (* constant 5056 *)
25226 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)).
25227 exact (fun (x: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)))).
25230 (* constant 5057 *)
25231 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)).
25232 exact (fun (x: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))).
25235 (* constant 5058 *)
25236 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)).
25237 exact (fun (x: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)))).
25240 (* constant 5059 *)
25241 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)).
25242 exact (fun (x: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))).
25245 (* constant 5060 *)
25246 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)).
25247 exact (fun (x: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)))).
25250 (* constant 5061 *)
25251 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)).
25252 exact (fun (x: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))).
25255 (* constant 5062 *)
25256 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)).
25257 exact (fun (x: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)))).
25260 (* constant 5063 *)
25261 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)).
25262 exact (fun (x: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))).
25265 (* constant 5064 *)
25266 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)).
25267 exact (fun (x: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)))).
25270 (* constant 5065 *)
25271 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))))).
25272 exact (fun (x:l_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))))))).
25275 (* constant 5066 *)
25276 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)))))).
25277 exact (fun (x:l_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)))))).
25280 (* constant 5067 *)
25281 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)))))).
25282 exact (fun (x:l_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)))))).
25285 (* constant 5068 *)
25286 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)))))).
25287 exact (fun (x:l_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))))))).
25290 (* constant 5069 *)
25291 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)))).
25292 exact (fun (x:l_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))))).
25295 (* constant 5070 *)
25296 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)))).
25297 exact (fun (x:l_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)))))).
25300 (* constant 5071 *)
25301 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)))).
25302 exact (fun (x:l_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)))))).
25305 (* constant 5072 *)
25306 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)))).
25307 exact (fun (x:l_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))))).
25310 (* constant 5073 *)
25311 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))))).
25312 exact (fun (x:l_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))))).
25315 (* constant 5074 *)
25316 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))))).
25317 exact (fun (x:l_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)))))).
25320 (* constant 5075 *)
25321 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))))).
25322 exact (fun (x:l_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))))).
25325 (* constant 5076 *)
25326 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)))).
25327 exact (fun (x:l_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))))).
25330 (* constant 5077 *)
25331 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)))).
25332 exact (fun (x:l_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)))))).
25335 (* constant 5078 *)
25336 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)))).
25337 exact (fun (x:l_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)))).
25340 (* constant 5079 *)
25341 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)))).
25342 exact (fun (x:l_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))))).
25345 (* constant 5080 *)
25346 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)).
25347 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)))).
25350 (* constant 5081 *)
25351 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)).
25352 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))).
25355 (* constant 5082 *)
25356 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).
25357 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)).
25360 (* constant 5083 *)
25361 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)).
25362 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)).
25365 (* constant 5084 *)
25366 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).
25367 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)).
25370 (* constant 5085 *)
25371 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)).
25372 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)).
25375 (* constant 5086 *)
25376 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))).
25377 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)))).
25380 (* constant 5087 *)
25381 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))).
25382 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))).
25385 (* constant 5088 *)
25386 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)).
25387 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)).
25390 (* constant 5089 *)
25391 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))).
25392 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)).
25395 (* constant 5090 *)
25396 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)).
25397 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)).
25400 (* constant 5091 *)
25401 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)).
25402 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)).
25405 (* constant 5092 *)
25406 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)).
25407 exact (fun (x: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))).
25410 (* constant 5093 *)
25411 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)).
25412 exact (fun (x: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))).
25415 (* constant 5094 *)
25416 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)).
25417 exact (fun (x: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))).
25420 (* constant 5095 *)
25421 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)).
25422 exact (fun (x: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))).
25425 (* constant 5096 *)
25426 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))))).
25427 exact (fun (x: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)))).
25430 (* constant 5097 *)
25431 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))))).
25432 exact (fun (x: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)))).
25435 (* constant 5098 *)
25436 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)))).
25437 exact (fun (x: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)))))).
25440 (* constant 5099 *)
25441 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)))).
25442 exact (fun (x: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)))).
25445 (* constant 5100 *)
25446 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)))).
25447 exact (fun (x: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))).
25450 (* constant 5101 *)
25451 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))).
25452 exact (fun (x: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))).
25455 (* constant 5102 *)
25456 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))).
25457 exact (fun (x: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))).
25460 (* constant 5103 *)
25461 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)))).
25462 exact (fun (x: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))).
25465 (* constant 5104 *)
25466 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))).
25467 exact (fun (x: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)))).
25470 (* constant 5105 *)
25471 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)))).
25472 exact (fun (x: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))).
25475 (* constant 5106 *)
25476 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))).
25477 exact (fun (x:l_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)))).
25480 (* constant 5107 *)
25481 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))).
25482 exact (fun (x:l_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)))).
25485 (* constant 5108 *)
25486 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))).
25487 exact (fun (x:l_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)))).
25490 (* constant 5109 *)
25491 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))).
25492 exact (fun (x:l_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)))).
25495 (* constant 5110 *)
25496 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))).
25497 exact (fun (x:l_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)))).
25500 (* constant 5111 *)
25501 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))).
25502 exact (fun (x:l_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)))).
25505 (* constant 5112 *)
25506 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))).
25507 exact (fun (x:l_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)))).
25510 (* constant 5113 *)
25511 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))).
25512 exact (fun (x:l_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)))).
25515 (* constant 5114 *)
25516 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))))))).
25517 exact (fun (x:l_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))))))).
25520 (* constant 5115 *)
25521 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))))).
25522 exact (fun (x:l_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)))))).
25525 (* constant 5116 *)
25526 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)))))).
25527 exact (fun (x:l_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))))).
25530 (* constant 5117 *)
25531 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)))).
25532 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))))).
25535 (* constant 5118 *)
25536 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)))).
25537 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)))).
25540 (* constant 5119 *)
25541 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))))).
25542 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)))).
25545 (* constant 5120 *)
25546 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))).
25547 exact (fun (x:l_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))).
25550 (* constant 5121 *)
25551 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))).
25552 exact (fun (x:l_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))).
25555 (* constant 5122 *)
25556 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))).
25557 exact (fun (x:l_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))).
25560 (* constant 5123 *)
25561 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))).
25562 exact (fun (x:l_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))).
25565 (* constant 5124 *)
25566 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))).
25567 exact (fun (x:l_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))).
25570 (* constant 5125 *)
25571 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))).
25572 exact (fun (x:l_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))).
25575 (* constant 5126 *)
25576 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))).
25577 exact (fun (x:l_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))).
25580 (* constant 5127 *)
25581 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))).
25582 exact (fun (x:l_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))).
25585 (* constant 5128 *)
25586 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)))))).
25587 exact (fun (x:l_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))))))).
25590 (* constant 5129 *)
25591 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))))).
25592 exact (fun (x:l_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)))))).
25595 (* constant 5130 *)
25596 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))))))).
25597 exact (fun (x:l_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)))).
25600 (* constant 5131 *)
25601 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))))).
25602 exact (fun (x:l_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))))).
25605 (* constant 5132 *)
25606 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)))))).
25607 exact (fun (x:l_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))))).
25610 (* constant 5133 *)
25611 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)))))).
25612 exact (fun (x:l_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)))).
25615 (* constant 5134 *)
25616 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))))).
25617 exact (fun (x:l_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)))).
25620 (* constant 5135 *)
25621 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))))).
25622 exact (fun (x:l_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))).
25625 (* constant 5136 *)
25626 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))))).
25627 exact (fun (x:l_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))).
25630 (* constant 5137 *)
25631 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)))).
25632 exact (fun (x:l_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)))).
25635 (* constant 5138 *)
25636 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)))).
25637 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)))).
25640 (* constant 5139 *)
25641 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)))))).
25642 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))))).
25645 (* constant 5140 *)
25646 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)))))).
25647 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)))))).
25650 (* constant 5141 *)
25651 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))))).
25652 exact (fun (x:l_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)))))).
25655 (* constant 5142 *)
25656 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))))).
25657 exact (fun (x:l_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)))))).
25660 (* constant 5143 *)
25661 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))))).
25662 exact (fun (x:l_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))))).
25665 (* constant 5144 *)
25666 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))))).
25667 exact (fun (x:l_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))).
25670 (* constant 5145 *)
25671 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))))).
25672 exact (fun (x:l_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))))).
25675 (* constant 5146 *)
25676 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))))).
25677 exact (fun (x:l_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))).
25680 (* constant 5147 *)
25681 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)))).
25682 exact (fun (x:l_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)))).
25685 (* constant 5148 *)
25686 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))))).
25687 exact (fun (x:l_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)))).
25690 (* constant 5149 *)
25691 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))))).
25692 exact (fun (x:l_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))))).
25695 (* constant 5150 *)
25696 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))))).
25697 exact (fun (x:l_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))))).
25700 (* constant 5151 *)
25701 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))))).
25702 exact (fun (x:l_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))).
25705 (* constant 5152 *)
25706 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)))).
25707 exact (fun (x:l_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)))).
25710 (* constant 5153 *)
25711 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))))).
25712 exact (fun (x:l_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)))).
25715 (* constant 5154 *)
25716 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)))))))).
25717 exact (fun (x:l_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))))))).
25720 (* constant 5155 *)
25721 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))))))).
25722 exact (fun (x:l_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))))))))).
25725 (* constant 5156 *)
25726 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))))))).
25727 exact (fun (x:l_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))))))).
25730 (* constant 5157 *)
25731 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))))))).
25732 exact (fun (x:l_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)))))))).
25735 (* constant 5158 *)
25736 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))).
25737 exact (fun (x:l_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)))).
25740 (* constant 5159 *)
25741 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))).
25742 exact (fun (x:l_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))).
25745 (* constant 5160 *)
25746 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)))).
25747 exact (fun (x:l_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))))).
25750 (* constant 5161 *)
25751 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)))))).
25752 exact (fun (x:l_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))))).
25755 (* constant 5162 *)
25756 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))).
25757 exact (fun (x:l_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))))).
25760 (* constant 5163 *)
25761 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)))))).
25762 exact (fun (x:l_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))))).
25765 (* constant 5164 *)
25766 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))))))).
25767 exact (fun (x:l_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))))).
25770 (* constant 5165 *)
25771 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))).
25772 exact (fun (x:l_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)))))).
25775 (* constant 5166 *)
25776 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))).
25777 exact (fun (x:l_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))))).
25780 (* constant 5167 *)
25781 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))).
25782 exact (fun (x:l_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))))).
25785 (* constant 5168 *)
25786 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))).
25787 exact (fun (x:l_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)))).
25790 (* constant 5169 *)
25791 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)))).
25792 exact (fun (x:l_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)))).
25795 (* constant 5170 *)
25796 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)))).
25797 exact (fun (x:l_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)))).
25800 (* constant 5171 *)
25801 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))).
25802 exact (fun (x:l_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)))).
25805 (* constant 5172 *)
25806 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))).
25807 exact (fun (x:l_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)))).
25810 (* constant 5173 *)
25811 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))))).
25812 exact (fun (x:l_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)))).
25815 (* constant 5174 *)
25816 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))).
25817 exact (fun (x:l_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)))).
25820 (* constant 5175 *)
25821 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)))).
25822 exact (fun (x:l_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)))).
25825 (* constant 5176 *)
25826 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)))))).
25827 exact (fun (x:l_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)))))).
25830 (* constant 5177 *)
25831 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))))).
25832 exact (fun (x:l_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)))))).
25835 (* constant 5178 *)
25836 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)))))).
25837 exact (fun (x:l_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)))))).
25840 (* constant 5179 *)
25841 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))))).
25842 exact (fun (x:l_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)))))).
25845 (* constant 5180 *)
25846 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)))))).
25847 exact (fun (x:l_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))))).
25850 (* constant 5181 *)
25851 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))))))).
25852 exact (fun (x:l_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)))))))).
25855 (* constant 5182 *)
25856 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))))))))).
25857 exact (fun (x:l_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))))))))).
25860 (* constant 5183 *)
25861 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)).
25862 exact (fun (x: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))).
25865 (* constant 5184 *)
25866 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)).
25867 exact (fun (x: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))).
25870 (* constant 5185 *)
25871 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)).
25872 exact (fun (x: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))).
25875 (* constant 5186 *)
25876 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)))).
25877 exact (fun (x:l_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)))).
25880 (* constant 5187 *)
25881 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))).
25882 exact (fun (x:l_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)))).
25885 (* constant 5188 *)
25886 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))))).
25887 exact (fun (x:l_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)))).
25890 (* constant 5189 *)
25891 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))))).
25892 exact (fun (x:l_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)))))).
25895 (* constant 5190 *)
25896 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)))).
25897 exact (fun (x:l_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)))).
25900 (* constant 5191 *)
25901 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)))).
25902 exact (fun (x:l_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))))).
25905 (* constant 5192 *)
25906 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)))).
25907 exact (fun (x:l_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)))).
25910 (* constant 5193 *)
25911 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))))).
25912 exact (fun (x:l_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))))).
25915 (* constant 5194 *)
25916 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)))).
25917 exact (fun (x:l_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)))).
25920 (* constant 5195 *)
25921 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)))).
25922 exact (fun (x:l_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)))).
25925 (* constant 5196 *)
25926 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)))).
25927 exact (fun (x:l_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))))).
25930 (* constant 5197 *)
25931 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)))).
25932 exact (fun (x:l_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)))))).
25935 (* constant 5198 *)
25936 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)))).
25937 exact (fun (x:l_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)))).
25940 (* constant 5199 *)
25941 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))))).
25942 exact (fun (x:l_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))))).
25945 (* constant 5200 *)
25946 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)))))).
25947 exact (fun (x:l_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)))))).
25950 (* constant 5201 *)
25951 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))))).
25952 exact (fun (x:l_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)))))).
25955 (* constant 5202 *)
25956 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)))))).
25957 exact (fun (x:l_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))))).
25960 (* constant 5203 *)
25961 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)))))).
25962 exact (fun (x:l_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)))))).
25965 (* constant 5204 *)
25966 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))))).
25967 exact (fun (x:l_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)))))).
25970 (* constant 5205 *)
25971 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)))))).
25972 exact (fun (x:l_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))))).
25975 (* constant 5206 *)
25976 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))))).
25977 exact (fun (x:l_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)))))).
25980 (* constant 5207 *)
25981 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)))))).
25982 exact (fun (x:l_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)))))).
25985 (* constant 5208 *)
25986 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))))).
25987 exact (fun (x:l_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)))))).
25990 (* constant 5209 *)
25991 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)))))).
25992 exact (fun (x:l_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)))))).
25995 (* constant 5210 *)
25996 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))).
25997 exact (fun (x:l_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)))).
26000 (* constant 5211 *)
26001 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))).
26002 exact (fun (x:l_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)))).
26005 (* constant 5212 *)
26006 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)))).
26007 exact (fun (x:l_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)))))).
26010 (* constant 5213 *)
26011 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)))).
26012 exact (fun (x:l_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))))).
26015 (* constant 5214 *)
26016 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)))))).
26017 exact (fun (x:l_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)))))).
26020 (* constant 5215 *)
26021 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))))).
26022 exact (fun (x:l_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)))))).
26025 (* constant 5216 *)
26026 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))))))).
26027 exact (fun (x:l_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)))))).
26030 (* constant 5217 *)
26031 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)))))).
26032 exact (fun (x:l_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))))))).
26035 (* constant 5218 *)
26036 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))))).
26037 exact (fun (x:l_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))))).
26040 (* constant 5219 *)
26041 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))))).
26042 exact (fun (x:l_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)))))).
26045 (* constant 5220 *)
26046 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))))).
26047 exact (fun (x:l_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))))).
26050 (* constant 5221 *)
26051 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))))))).
26052 exact (fun (x:l_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))))))).
26055 (* constant 5222 *)
26056 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)))))).
26057 exact (fun (x:l_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)))))).
26060 (* constant 5223 *)
26061 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)))))).
26062 exact (fun (x:l_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)))))).
26065 (* constant 5224 *)
26066 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)))))).
26067 exact (fun (x:l_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)))))).
26070 (* constant 5225 *)
26071 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)))))).
26072 exact (fun (x:l_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)))))))).
26075 (* constant 5226 *)
26076 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)))))).
26077 exact (fun (x:l_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)))))).
26080 (* constant 5227 *)
26081 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))))))).
26082 exact (fun (x:l_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))))))).
26085 (* constant 5228 *)
26086 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)))))))).
26087 exact (fun (x:l_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)))))))).
26090 (* constant 5229 *)
26091 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))))))).
26092 exact (fun (x:l_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)))))))).
26095 (* constant 5230 *)
26096 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)))))))).
26097 exact (fun (x:l_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))))))).
26100 (* constant 5231 *)
26101 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))))))))).
26102 exact (fun (x:l_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))))))))).
26105 (* constant 5232 *)
26106 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)))))))).
26107 exact (fun (x:l_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))))))))).
26110 (* constant 5233 *)
26111 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))))))))).
26112 exact (fun (x:l_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)))))))).
26115 (* constant 5234 *)
26116 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)).
26117 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)))).
26120 (* constant 5235 *)
26121 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)).
26122 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))).
26125 (* constant 5236 *)
26126 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)))).
26127 exact (fun (x:l_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)))))).
26130 (* constant 5237 *)
26131 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)))).
26132 exact (fun (x:l_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))))).
26135 (* constant 5238 *)
26136 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)))))))).
26137 exact (fun (x:l_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))))))))).
26140 (* constant 5239 *)
26141 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)))))))).
26142 exact (fun (x:l_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)))))))))).
26145 (* constant 5240 *)
26146 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))))))))).
26147 exact (fun (x:l_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))))))))).
26150 (* constant 5241 *)
26151 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)))))))).
26152 exact (fun (x:l_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))))))))).
26155 (* constant 5242 *)
26156 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))))))))).
26157 exact (fun (x:l_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))))))))).
26160 (* constant 5243 *)
26161 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)))))))).
26162 exact (fun (x:l_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)))))))).
26165 (* constant 5244 *)
26166 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)))))))).
26167 exact (fun (x:l_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))))))))).
26170 (* constant 5245 *)
26171 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)))))))).
26172 exact (fun (x:l_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)))))))).
26175 (* constant 5246 *)
26176 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))))))))).
26177 exact (fun (x:l_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)))))))))).
26180 (* constant 5247 *)
26181 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)))))))).
26182 exact (fun (x:l_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))))))))).
26185 (* constant 5248 *)
26186 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))))))))).
26187 exact (fun (x:l_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))))))))).
26190 (* constant 5249 *)
26191 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)))))))).
26192 exact (fun (x:l_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)))))))).
26195 (* constant 5250 *)
26196 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))))).
26197 exact (fun (x:l_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))))))).
26200 (* constant 5251 *)
26201 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)))))).
26202 exact (fun (x:l_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))))).
26205 (* constant 5252 *)
26206 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))))).
26207 exact (fun (x:l_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)))).
26210 (* constant 5253 *)
26211 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)))))))).
26212 exact (fun (x:l_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)))))))).
26215 (* constant 5254 *)
26216 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))))).
26217 exact (fun (x:l_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))))))).
26220 (* constant 5255 *)
26221 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)))))).
26222 exact (fun (x:l_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))))).
26225 (* constant 5256 *)
26226 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))))).
26227 exact (fun (x:l_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)))).
26230 (* constant 5257 *)
26231 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)))))))).
26232 exact (fun (x:l_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)))))))).
26235 (* constant 5258 *)
26236 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).
26237 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))).
26240 (* constant 5259 *)
26241 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)))).
26242 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)))).
26245 (* constant 5260 *)
26246 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)))).
26247 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))).
26250 (* constant 5261 *)
26251 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)))).
26252 exact (fun (x:l_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))).
26255 (* constant 5262 *)
26256 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).
26257 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)).
26260 (* constant 5263 *)
26261 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)).
26262 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))))).
26265 (* constant 5264 *)
26266 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)).
26267 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))).
26270 (* constant 5265 *)
26271 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)).
26272 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)))).
26275 (* constant 5266 *)
26276 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)).
26277 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)))).
26280 (* constant 5267 *)
26281 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)).
26282 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))).
26285 (* constant 5268 *)
26286 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)).
26287 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))).
26290 (* constant 5269 *)
26291 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))).
26292 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))).
26295 (* constant 5270 *)
26296 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)).
26297 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)))).
26300 (* constant 5271 *)
26301 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)).
26302 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))).
26305 (* constant 5272 *)
26306 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)).
26307 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))).
26310 (* constant 5273 *)
26311 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))).
26312 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))).
26315 (* constant 5274 *)
26316 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)))).
26317 exact (fun (x: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))))).
26320 (* constant 5275 *)
26321 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)))).
26322 exact (fun (x: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))).
26325 (* constant 5276 *)
26326 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))))).
26327 exact (fun (x: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))))).
26330 (* constant 5277 *)
26331 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)))).
26332 exact (fun (x: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))))).
26335 (* constant 5278 *)
26336 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)))).
26337 exact (fun (x: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))).
26340 (* constant 5279 *)
26341 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))).
26342 exact (fun (x: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))).
26345 (* constant 5280 *)
26346 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)))).
26347 exact (fun (x: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))).
26350 (* constant 5281 *)
26351 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)))).
26352 exact (fun (x: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)))).
26355 (* constant 5282 *)
26356 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)))).
26357 exact (fun (x: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))).
26360 (* constant 5283 *)
26361 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)))).
26362 exact (fun (x:l_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))).
26365 (* constant 5284 *)
26366 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))))).
26367 exact (fun (x:l_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)))).
26370 (* constant 5285 *)
26371 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))))).
26372 exact (fun (x:l_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))).
26375 (* constant 5286 *)
26376 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))))).
26377 exact (fun (x:l_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)))).
26380 (* constant 5287 *)
26381 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))).
26382 exact (fun (x:l_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))).
26385 (* constant 5288 *)
26386 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))))).
26387 exact (fun (x:l_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))))).
26390 (* constant 5289 *)
26391 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))))).
26392 exact (fun (x:l_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)))).
26395 (* constant 5290 *)
26396 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).
26397 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)).
26400 (* constant 5291 *)
26401 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)))).
26402 exact (fun (x:l_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))).
26405 (* constant 5292 *)
26406 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))).
26407 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)))).
26410 (* constant 5293 *)
26411 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)).
26412 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))).
26415 (* constant 5294 *)
26416 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))).
26417 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)).
26420 (* constant 5295 *)
26421 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).
26422 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))).
26425 (* constant 5296 *)
26426 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))).
26427 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))).
26430 (* constant 5297 *)
26431 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)))).
26432 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))).
26435 (* constant 5298 *)
26436 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)).
26437 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))).
26440 (* constant 5299 *)
26441 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)))).
26442 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)).
26445 (* constant 5300 *)
26446 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)))).
26447 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))).
26450 (* constant 5301 *)
26451 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)))).
26452 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))))).
26455 (* constant 5302 *)
26456 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)))).
26457 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)).
26460 (* constant 5303 *)
26461 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))))).
26462 exact (fun (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))))).
26465 (* constant 5304 *)
26466 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))))).
26467 exact (fun (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)))))).
26470 (* constant 5305 *)
26471 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))))))).
26472 exact (fun (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)))))))).
26475 (* constant 5306 *)
26476 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))))))).
26477 exact (fun (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)))))))))).
26480 (* constant 5307 *)
26481 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)))))).
26482 exact (fun (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))))))).
26485 (* constant 5308 *)
26486 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)))).
26487 exact (fun (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)))))).
26490 (* constant 5309 *)
26491 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))).
26492 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)).
26495 (* constant 5310 *)
26496 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))).
26497 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)).
26500 (* constant 5311 *)
26501 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)).
26502 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)))).
26505 (* constant 5312 *)
26506 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).
26507 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))).
26510 (* constant 5313 *)
26511 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)).
26512 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))).
26515 (* constant 5314 *)
26516 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)))))).
26517 exact (fun (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)))))).
26520 (* constant 5315 *)
26521 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)))))).
26522 exact (fun (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)))))).
26525 (* constant 5316 *)
26526 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))))))).
26527 exact (fun (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)))))).
26530 (* constant 5317 *)
26531 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))))))).
26532 exact (fun (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))))))).
26535 (* constant 5318 *)
26536 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))))).
26537 exact (fun (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)))))).
26540 (* constant 5319 *)
26541 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)).
26542 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)))).
26545 (* constant 5320 *)
26546 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)).
26547 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)))).
26550 (* constant 5321 *)
26551 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).
26552 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)))).
26555 (* constant 5322 *)
26556 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)).
26557 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))).
26560 (* constant 5323 *)
26561 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))).
26562 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)).
26565 (* constant 5324 *)
26566 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))).
26567 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)).
26570 (* constant 5325 *)
26571 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))))).
26572 exact (fun (x:l_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)))).
26575 (* constant 5326 *)
26576 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)))))).
26577 exact (fun (x:l_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)))))).
26580 (* constant 5327 *)
26581 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))))).
26582 exact (fun (x: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)))).
26585 (* constant 5328 *)
26586 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))))).
26587 exact (fun (x: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)))).
26590 (* constant 5329 *)
26591 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))).
26592 exact (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))))).
26595 (* constant 5330 *)
26596 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)).
26597 exact (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)))).
26600 (* constant 5331 *)
26601 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))).
26602 exact (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)))).
26605 (* constant 5332 *)
26606 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)))).
26607 exact (fun (x: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))))).
26610 (* constant 5333 *)
26611 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))).
26612 exact (fun (x:l_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))))).
26615 (* constant 5334 *)
26616 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)))).
26617 exact (fun (x: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))).
26620 (* constant 5335 *)
26621 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)))).
26622 exact (fun (x: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))).
26625 (* constant 5336 *)
26626 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))).
26627 exact (fun (x:l_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)))).
26630 (* constant 5337 *)
26631 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)))).
26632 exact (fun (x:l_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))))).
26635 (* constant 5338 *)
26636 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))))).
26637 exact (fun (x:l_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))))).
26640 (* constant 5339 *)
26641 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)))))).
26642 exact (fun (x:l_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))).
26645 (* constant 5340 *)
26646 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).
26647 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))).
26650 (* constant 5341 *)
26651 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)).
26652 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))).
26655 (* constant 5342 *)
26656 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))).
26657 exact (fun (x:l_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)))).
26660 (* constant 5343 *)
26661 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))).
26662 exact (fun (x:l_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))))).
26665 (* constant 5344 *)
26666 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))).
26667 exact (fun (x:l_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))).
26670 (* constant 5345 *)
26671 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))))).
26672 exact (fun (x:l_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))))).
26675 (* constant 5346 *)
26676 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))).
26677 exact (fun (x:l_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)))).
26680 (* constant 5347 *)
26681 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))).
26682 exact (fun (x:l_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)))).
26685 (* constant 5348 *)
26686 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))).
26687 exact (fun (x:l_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)))).
26690 (* constant 5349 *)
26691 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))).
26692 exact (fun (x:l_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)))).
26695 (* constant 5350 *)
26696 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))).
26697 exact (fun (x:l_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)))).
26700 (* constant 5351 *)
26701 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))).
26702 exact (fun (x:l_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)))).
26705 (* constant 5352 *)
26706 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))).
26707 exact (fun (x:l_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))))).
26710 (* constant 5353 *)
26711 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))).
26712 exact (fun (x:l_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))))).
26715 (* constant 5354 *)
26716 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)))).
26717 exact (fun (x:l_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))))).
26720 (* constant 5355 *)
26721 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))))).
26722 exact (fun (x:l_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)))))).
26725 (* constant 5356 *)
26726 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))))).
26727 exact (fun (x:l_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))))).
26730 (* constant 5357 *)
26731 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))))).
26732 exact (fun (x:l_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)))).
26735 (* constant 5358 *)
26736 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)))).
26737 exact (fun (x: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))).
26740 (* constant 5359 *)
26741 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)))).
26742 exact (fun (x: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))).
26745 (* constant 5360 *)
26746 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)))).
26747 exact (fun (x: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)).
26750 (* constant 5361 *)
26751 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))).
26752 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)))).
26755 (* constant 5362 *)
26756 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))).
26757 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))).
26760 (* constant 5363 *)
26761 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))).
26762 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)))).
26765 (* constant 5364 *)
26766 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))).
26767 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))).
26770 (* constant 5365 *)
26771 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)).
26772 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)).
26775 (* constant 5366 *)
26776 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)).
26777 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)).
26780 (* constant 5367 *)
26781 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))).
26782 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)).
26785 (* constant 5368 *)
26786 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)).
26787 exact (fun (x: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)))).
26790 (* constant 5369 *)
26791 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))).
26792 exact (fun (x: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)))).
26795 (* constant 5370 *)
26796 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)))).
26797 exact (fun (x: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))).
26800 (* constant 5371 *)
26801 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)))).
26802 exact (fun (x: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)))).
26805 (* constant 5372 *)
26806 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)))).
26807 exact (fun (x: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)))).
26810 (* constant 5373 *)
26811 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))))).
26812 exact (fun (x: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))).
26815 (* constant 5374 *)
26816 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)).
26817 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)).
26820 (* constant 5375 *)
26821 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))))).
26822 exact (fun (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))))).
26825 (* constant 5376 *)
26826 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))))).
26827 exact (fun (x: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))).
26830 (* constant 5377 *)
26831 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)).
26832 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))).
26835 (* constant 5378 *)
26836 Definition l_e_st_eq_landau_n_8274_prop2 : (forall (x:l_e_st_eq_landau_n_nat), Prop).
26837 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)))).
26840 (* constant 5379 *)
26841 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))).
26842 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))).
26845 (* constant 5380 *)
26846 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))).
26847 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))).
26850 (* constant 5381 *)
26851 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)))).
26852 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)))).
26855 (* constant 5382 *)
26856 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))).
26857 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))).
26860 (* constant 5383 *)
26861 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))))).
26862 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)))).
26865 (* constant 5384 *)
26866 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))))).
26867 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))))).
26870 (* constant 5385 *)
26871 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))))))).
26872 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)))))).
26875 (* constant 5386 *)
26876 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)))))).
26877 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)))))).
26880 (* constant 5387 *)
26881 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))))).
26882 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))))).
26885 (* constant 5388 *)
26886 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))))))).
26887 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)))))).
26890 (* constant 5389 *)
26891 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)))))).
26892 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)))))).
26895 (* constant 5390 *)
26896 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))))).
26897 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))))).
26900 (* constant 5391 *)
26901 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)))).
26902 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)))).
26905 (* constant 5392 *)
26906 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)))).
26907 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)))).
26910 (* constant 5393 *)
26911 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))).
26912 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))).
26915 (* constant 5394 *)
26916 Definition l_e_st_eq_landau_n_8274_t14 : l_e_st_eq_landau_n_8274_prop2 l_e_st_eq_landau_n_1.
26917 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)).
26920 (* constant 5395 *)
26921 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).
26922 exact (fun (x:l_e_st_eq_landau_n_nat) => l_e_st_eq_landau_n_suc x).
26925 (* constant 5396 *)
26926 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))))).
26927 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))))).
26930 (* constant 5397 *)
26931 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)))).
26932 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)))).
26935 (* constant 5398 *)
26936 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)))).
26937 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)))).
26940 (* constant 5399 *)
26941 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)))).
26942 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))))).
26945 (* constant 5400 *)
26946 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))))).
26947 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)))).
26950 (* constant 5401 *)
26951 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))))).
26952 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))))).
26955 (* constant 5402 *)
26956 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)))).
26957 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))))).
26960 (* constant 5403 *)
26961 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)))))).
26962 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)))))).
26965 (* constant 5404 *)
26966 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)))))).
26967 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)))))).
26970 (* constant 5405 *)
26971 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)))))))).
26972 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)))))))).
26975 (* constant 5406 *)
26976 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))))))))).
26977 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))))))))).
26980 (* constant 5407 *)
26981 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))))))))).
26982 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))))))))).
26985 (* constant 5408 *)
26986 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))))))))).
26987 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))))))))).
26990 (* constant 5409 *)
26991 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))))))))).
26992 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))))))))).
26995 (* constant 5410 *)
26996 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)))))))))).
26997 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))))))))).
27000 (* constant 5411 *)
27001 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))))))))))).
27002 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))))))))).
27005 (* constant 5412 *)
27006 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)))))))))).
27007 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)))))))))).
27010 (* constant 5413 *)
27011 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)))))))))).
27012 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))))))))).
27015 (* constant 5414 *)
27016 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)))))))).
27017 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)))))))))).
27020 (* constant 5415 *)
27021 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)))))))))).
27022 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)))))))))).
27025 (* constant 5416 *)
27026 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)))))))).
27027 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))))))))).
27030 (* constant 5417 *)
27031 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)))))))).
27032 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))))))))).
27035 (* constant 5418 *)
27036 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))))))))).
27037 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)))))))))).
27040 (* constant 5419 *)
27041 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))))))))).
27042 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))))))))).
27045 (* constant 5420 *)
27046 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))))))))).
27047 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))))))))).
27050 (* constant 5421 *)
27051 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))))))))).
27052 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)))))))).
27055 (* constant 5422 *)
27056 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))))))))))).
27057 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)))))))))).
27060 (* constant 5423 *)
27061 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)))))))))))).
27062 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))))))))))).
27065 (* constant 5424 *)
27066 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))))))))))).
27067 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))))))))))).
27070 (* constant 5425 *)
27071 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))))))))))).
27072 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))))))))))).
27075 (* constant 5426 *)
27076 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)))))))))).
27077 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))))))))))).
27080 (* constant 5427 *)
27081 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)))))))).
27082 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)))))))).
27085 (* constant 5428 *)
27086 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)))))))).
27087 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))))))))).
27090 (* constant 5429 *)
27091 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)))))))).
27092 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))))))))).
27095 (* constant 5430 *)
27096 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)))))))).
27097 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))))))))).
27100 (* constant 5431 *)
27101 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)))))))).
27102 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))))))))).
27105 (* constant 5432 *)
27106 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))))))))).
27107 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))))))))).
27110 (* constant 5433 *)
27111 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)))))))))).
27112 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))))))))).
27115 (* constant 5434 *)
27116 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))))))))))).
27117 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))))))))).
27120 (* constant 5435 *)
27121 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)))))))))).
27122 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))))))))).
27125 (* constant 5436 *)
27126 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))))))))).
27127 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)))))))))).
27130 (* constant 5437 *)
27131 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)))))))))).
27132 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))))))))).
27135 (* constant 5438 *)
27136 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)))))))).
27137 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))))))))).
27140 (* constant 5439 *)
27141 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)))))))))).
27142 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)))))))))).
27145 (* constant 5440 *)
27146 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))))))))).
27147 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))))))))).
27150 (* constant 5441 *)
27151 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))))))))).
27152 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))))))))).
27155 (* constant 5442 *)
27156 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)))))))).
27157 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))))))))).
27160 (* constant 5443 *)
27161 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)))))))).
27162 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))))))))).
27165 (* constant 5444 *)
27166 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)))))))))).
27167 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))))))))).
27170 (* constant 5445 *)
27171 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)))))))))).
27172 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)))))))))).
27175 (* constant 5446 *)
27176 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))))))))))).
27177 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))))))))).
27180 (* constant 5447 *)
27181 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))))))))))).
27182 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))))))))).
27185 (* constant 5448 *)
27186 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)))))))))).
27187 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)))))))))).
27190 (* constant 5449 *)
27191 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)))))))))).
27192 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)))))))))).
27195 (* constant 5450 *)
27196 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)))))))).
27197 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))))))))).
27200 (* constant 5451 *)
27201 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)))))))).
27202 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)))))))).
27205 (* constant 5452 *)
27206 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)))))))).
27207 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)))))))).
27210 (* constant 5453 *)
27211 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))))))).
27212 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)))))))).
27215 (* constant 5454 *)
27216 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)))))))).
27217 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)))))))).
27220 (* constant 5455 *)
27221 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))))))))).
27222 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)))))))).
27225 (* constant 5456 *)
27226 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)))))))).
27227 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)))))))).
27230 (* constant 5457 *)
27231 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))))))))).
27232 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))))))))).
27235 (* constant 5458 *)
27236 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)))))))).
27237 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)))))))).
27240 (* constant 5459 *)
27241 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)))))))).
27242 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))))))).
27245 (* constant 5460 *)
27246 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))))))).
27247 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)))))))).
27250 (* constant 5461 *)
27251 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)))))).
27252 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))))))).
27255 (* constant 5462 *)
27256 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))))).
27257 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)))))).
27260 (* constant 5463 *)
27261 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))).
27262 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)))).
27265 (* constant 5464 *)
27266 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).
27267 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).
27270 (* constant 5465 *)
27271 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))))).
27272 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))).
27275 (* constant 5466 *)
27276 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))))).
27277 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)))).
27280 (* constant 5467 *)
27281 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)).
27282 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)).
27285 (* constant 5468 *)
27286 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))))).
27287 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27290 (* constant 5469 *)
27291 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)))).
27292 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27295 (* constant 5470 *)
27296 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)))).
27297 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27300 (* constant 5471 *)
27301 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)))).
27302 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))))).
27305 (* constant 5472 *)
27306 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)).
27307 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)).
27310 (* constant 5473 *)
27311 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)).
27312 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)).
27315 (* constant 5474 *)
27316 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)))).
27317 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27320 (* constant 5475 *)
27321 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)))).
27322 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27325 (* constant 5476 *)
27326 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)))).
27327 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27330 (* constant 5477 *)
27331 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)))).
27332 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27335 (* constant 5478 *)
27336 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)))))).
27337 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27340 (* constant 5479 *)
27341 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))))))).
27342 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27345 (* constant 5480 *)
27346 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))))))).
27347 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27350 (* constant 5481 *)
27351 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)))).
27352 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27355 (* constant 5482 *)
27356 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)))).
27357 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27360 (* constant 5483 *)
27361 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)))).
27362 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27365 (* constant 5484 *)
27366 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))))).
27367 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27370 (* constant 5485 *)
27371 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)))))))))).
27372 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27375 (* constant 5486 *)
27376 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)))))))))).
27377 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27380 (* constant 5487 *)
27381 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))))))))))).
27382 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27385 (* constant 5488 *)
27386 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))))))).
27387 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27390 (* constant 5489 *)
27391 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)))))).
27392 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27395 (* constant 5490 *)
27396 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)))))).
27397 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27400 (* constant 5491 *)
27401 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))))))))).
27402 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27405 (* constant 5492 *)
27406 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)))))))))).
27407 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27410 (* constant 5493 *)
27411 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)))))))))).
27412 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27415 (* constant 5494 *)
27416 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)))))))).
27417 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27420 (* constant 5495 *)
27421 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))))))).
27422 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27425 (* constant 5496 *)
27426 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))))))).
27427 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27430 (* constant 5497 *)
27431 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))))))))))).
27432 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))))).
27435 (* constant 5498 *)
27436 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)))))))))).
27437 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))))).
27440 (* constant 5499 *)
27441 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))))))))))).
27442 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))))).
27445 (* constant 5500 *)
27446 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)))))))))))))).
27447 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))))).
27450 (* constant 5501 *)
27451 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)))))))))))).
27452 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))))).
27455 (* constant 5502 *)
27456 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)))))))))).
27457 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))))).
27460 (* constant 5503 *)
27461 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))))))))))).
27462 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))))).
27465 (* constant 5504 *)
27466 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))))))))))).
27467 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))))).
27470 (* constant 5505 *)
27471 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)))))))))).
27472 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))))).
27475 (* constant 5506 *)
27476 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))))))))))).
27477 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))))).
27480 (* constant 5507 *)
27481 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)))))))))).
27482 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))))).
27485 (* constant 5508 *)
27486 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)))))))).
27487 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27490 (* constant 5509 *)
27491 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))))))))))).
27492 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27495 (* constant 5510 *)
27496 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))))))))).
27497 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27500 (* constant 5511 *)
27501 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))))))))).
27502 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27505 (* constant 5512 *)
27506 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))))))))).
27507 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27510 (* constant 5513 *)
27511 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))))))))).
27512 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27515 (* constant 5514 *)
27516 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))))))))).
27517 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27520 (* constant 5515 *)
27521 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))))))).
27522 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27525 (* constant 5516 *)
27526 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))).
27527 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27530 (* constant 5517 *)
27531 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)).
27532 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
27535 (* constant 5518 *)
27536 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)).
27537 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27540 (* constant 5519 *)
27541 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)))).
27542 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27545 (* constant 5520 *)
27546 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))))))).
27547 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27550 (* constant 5521 *)
27551 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)).
27552 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27555 (* constant 5522 *)
27556 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)).
27557 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
27560 (* constant 5523 *)
27561 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).
27562 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)).
27565 (* constant 5524 *)
27566 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))).
27567 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
27570 (* constant 5525 *)
27571 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))))).
27572 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27575 (* constant 5526 *)
27576 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))))).
27577 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27580 (* constant 5527 *)
27581 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))))).
27582 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27585 (* constant 5528 *)
27586 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))))).
27587 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27590 (* constant 5529 *)
27591 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))))).
27592 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27595 (* constant 5530 *)
27596 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)))))).
27597 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27600 (* constant 5531 *)
27601 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)))))).
27602 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27605 (* constant 5532 *)
27606 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)))))).
27607 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27610 (* constant 5533 *)
27611 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)))))))).
27612 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27615 (* constant 5534 *)
27616 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)))))))).
27617 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27620 (* constant 5535 *)
27621 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)))).
27622 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27625 (* constant 5536 *)
27626 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))))).
27627 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27630 (* constant 5537 *)
27631 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))))))).
27632 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27635 (* constant 5538 *)
27636 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))))))).
27637 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27640 (* constant 5539 *)
27641 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))))).
27642 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27645 (* constant 5540 *)
27646 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))))))).
27647 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27650 (* constant 5541 *)
27651 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)))))).
27652 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27655 (* constant 5542 *)
27656 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))))))).
27657 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27660 (* constant 5543 *)
27661 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)))))).
27662 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27665 (* constant 5544 *)
27666 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)))))).
27667 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27670 (* constant 5545 *)
27671 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)))))).
27672 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27675 (* constant 5546 *)
27676 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)))))).
27677 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27680 (* constant 5547 *)
27681 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)))))).
27682 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27685 (* constant 5548 *)
27686 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)))))).
27687 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27690 (* constant 5549 *)
27691 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)))))).
27692 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27695 (* constant 5550 *)
27696 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))))).
27697 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27700 (* constant 5551 *)
27701 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)))))))).
27702 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27705 (* constant 5552 *)
27706 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))))))).
27707 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27710 (* constant 5553 *)
27711 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)))))))).
27712 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27715 (* constant 5554 *)
27716 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)))))))).
27717 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27720 (* constant 5555 *)
27721 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)))))))).
27722 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27725 (* constant 5556 *)
27726 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))))))))))).
27727 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27730 (* constant 5557 *)
27731 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))))))))).
27732 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27735 (* constant 5558 *)
27736 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)))))))))).
27737 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27740 (* constant 5559 *)
27741 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)))))))))).
27742 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27745 (* constant 5560 *)
27746 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)))))))).
27747 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27750 (* constant 5561 *)
27751 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))))))))).
27752 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27755 (* constant 5562 *)
27756 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)))))))))).
27757 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27760 (* constant 5563 *)
27761 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))))))))).
27762 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27765 (* constant 5564 *)
27766 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))))))))).
27767 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27770 (* constant 5565 *)
27771 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)))))))))).
27772 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27775 (* constant 5566 *)
27776 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))))))))).
27777 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27780 (* constant 5567 *)
27781 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)))))))))).
27782 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27785 (* constant 5568 *)
27786 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))))))))).
27787 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27790 (* constant 5569 *)
27791 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)))))))))).
27792 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27795 (* constant 5570 *)
27796 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)))))))))).
27797 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27800 (* constant 5571 *)
27801 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)))))))))).
27802 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27805 (* constant 5572 *)
27806 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)))))))))).
27807 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27810 (* constant 5573 *)
27811 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)))))))))).
27812 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27815 (* constant 5574 *)
27816 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))))))))))).
27817 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27820 (* constant 5575 *)
27821 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)))))))))).
27822 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27825 (* constant 5576 *)
27826 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))))))))))).
27827 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27830 (* constant 5577 *)
27831 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))))))))))).
27832 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27835 (* constant 5578 *)
27836 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)))))))))).
27837 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27840 (* constant 5579 *)
27841 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))))))))).
27842 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27845 (* constant 5580 *)
27846 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))))).
27847 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27850 (* constant 5581 *)
27851 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))))).
27852 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27855 (* constant 5582 *)
27856 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)))).
27857 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27860 (* constant 5583 *)
27861 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)))).
27862 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27865 (* constant 5584 *)
27866 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)).
27867 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)).
27870 (* constant 5585 *)
27871 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))).
27872 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
27875 (* constant 5586 *)
27876 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)))).
27877 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27880 (* constant 5587 *)
27881 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)))).
27882 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
27885 (* constant 5588 *)
27886 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)))).
27887 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27890 (* constant 5589 *)
27891 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)))).
27892 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27895 (* constant 5590 *)
27896 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)))).
27897 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27900 (* constant 5591 *)
27901 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))))).
27902 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
27905 (* constant 5592 *)
27906 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))))).
27907 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27910 (* constant 5593 *)
27911 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)))))))).
27912 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27915 (* constant 5594 *)
27916 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)))))).
27917 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27920 (* constant 5595 *)
27921 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))))))).
27922 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27925 (* constant 5596 *)
27926 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)))))).
27927 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27930 (* constant 5597 *)
27931 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)))))).
27932 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
27935 (* constant 5598 *)
27936 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)))))).
27937 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27940 (* constant 5599 *)
27941 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))))))).
27942 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27945 (* constant 5600 *)
27946 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))))))).
27947 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27950 (* constant 5601 *)
27951 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)))))).
27952 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
27955 (* constant 5602 *)
27956 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)))))))).
27957 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))))).
27960 (* constant 5603 *)
27961 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))))))).
27962 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27965 (* constant 5604 *)
27966 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))))))).
27967 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
27970 (* constant 5605 *)
27971 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))))))).
27972 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27975 (* constant 5606 *)
27976 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))))))))).
27977 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27980 (* constant 5607 *)
27981 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))))))))))).
27982 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27985 (* constant 5608 *)
27986 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))))))))).
27987 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27990 (* constant 5609 *)
27991 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)))))))))).
27992 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
27995 (* constant 5610 *)
27996 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))))))))).
27997 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
28000 (* constant 5611 *)
28001 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))))))))).
28002 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
28005 (* constant 5612 *)
28006 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))))))))).
28007 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
28010 (* constant 5613 *)
28011 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)))))))))).
28012 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))))).
28015 (* constant 5614 *)
28016 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)))))).
28017 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))))).
28020 (* constant 5615 *)
28021 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)))))).
28022 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
28025 (* constant 5616 *)
28026 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))))))).
28027 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
28030 (* constant 5617 *)
28031 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))).
28032 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28035 (* constant 5618 *)
28036 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))).
28037 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28040 (* constant 5619 *)
28041 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)))).
28042 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28045 (* constant 5620 *)
28046 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)))).
28047 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28050 (* constant 5621 *)
28051 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)))).
28052 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28055 (* constant 5622 *)
28056 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)))).
28057 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28060 (* constant 5623 *)
28061 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)))).
28062 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28065 (* constant 5624 *)
28066 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)))).
28067 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28070 (* constant 5625 *)
28071 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)))).
28072 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28075 (* constant 5626 *)
28076 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)))).
28077 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28080 (* constant 5627 *)
28081 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)))).
28082 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28085 (* constant 5628 *)
28086 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))))))).
28087 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
28090 (* constant 5629 *)
28091 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))))).
28092 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
28095 (* constant 5630 *)
28096 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))))).
28097 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28100 (* constant 5631 *)
28101 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)))).
28102 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28105 (* constant 5632 *)
28106 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)))))).
28107 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
28110 (* constant 5633 *)
28111 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))))))).
28112 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28115 (* constant 5634 *)
28116 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)))))).
28117 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
28120 (* constant 5635 *)
28121 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)))))).
28122 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
28125 (* constant 5636 *)
28126 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))))).
28127 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
28130 (* constant 5637 *)
28131 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)))).
28132 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28135 (* constant 5638 *)
28136 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))))).
28137 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28140 (* constant 5639 *)
28141 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))))).
28142 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
28145 (* constant 5640 *)
28146 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))))).
28147 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28150 (* constant 5641 *)
28151 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))))).
28152 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28155 (* constant 5642 *)
28156 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)))))).
28157 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
28160 (* constant 5643 *)
28161 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)))))).
28162 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28165 (* constant 5644 *)
28166 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)))))).
28167 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28170 (* constant 5645 *)
28171 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))))))).
28172 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28175 (* constant 5646 *)
28176 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))))))).
28177 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28180 (* constant 5647 *)
28181 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))))))).
28182 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28185 (* constant 5648 *)
28186 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))).
28187 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28190 (* constant 5649 *)
28191 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)).
28192 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)).
28195 (* constant 5650 *)
28196 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)).
28197 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)).
28200 (* constant 5651 *)
28201 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))).
28202 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28205 (* constant 5652 *)
28206 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)))).
28207 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28210 (* constant 5653 *)
28211 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))))))).
28212 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28215 (* constant 5654 *)
28216 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))))).
28217 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
28220 (* constant 5655 *)
28221 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)))))).
28222 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
28225 (* constant 5656 *)
28226 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))))))).
28227 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
28230 (* constant 5657 *)
28231 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))))).
28232 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))))).
28235 (* constant 5658 *)
28236 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)))))).
28237 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
28240 (* constant 5659 *)
28241 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)))))).
28242 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
28245 (* constant 5660 *)
28246 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)))))).
28247 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
28250 (* constant 5661 *)
28251 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)).
28252 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)).
28255 (* constant 5662 *)
28256 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)).
28257 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)))).
28260 (* constant 5663 *)
28261 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).
28262 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)).
28265 (* constant 5664 *)
28266 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).
28267 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)).
28270 (* constant 5665 *)
28271 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)))).
28272 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)))).
28275 (* constant 5666 *)
28276 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))))).
28277 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)))).
28280 (* constant 5667 *)
28281 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))))).
28282 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))).
28285 (* constant 5668 *)
28286 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))))).
28287 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))).
28290 (* constant 5669 *)
28291 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)))).
28292 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))))).
28295 (* constant 5670 *)
28296 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)))).
28297 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)))).
28300 (* constant 5671 *)
28301 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))))).
28302 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)))).
28305 (* constant 5672 *)
28306 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)))).
28307 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)))).
28310 (* constant 5673 *)
28311 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)))).
28312 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)))).
28315 (* constant 5674 *)
28316 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)))).
28317 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)).
28320 (* constant 5675 *)
28321 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)).
28322 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28325 (* constant 5676 *)
28326 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))).
28327 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)).
28330 (* constant 5677 *)
28331 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))))).
28332 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)).
28335 (* constant 5678 *)
28336 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)))).
28337 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28340 (* constant 5679 *)
28341 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)))).
28342 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28345 (* constant 5680 *)
28346 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)))).
28347 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28350 (* constant 5681 *)
28351 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)))).
28352 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28355 (* constant 5682 *)
28356 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)))).
28357 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28360 (* constant 5683 *)
28361 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))))).
28362 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28365 (* constant 5684 *)
28366 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))))).
28367 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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))).
28370 (* constant 5685 *)
28371 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).
28372 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))))).
28375 (* constant 5686 *)
28376 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)).
28377 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))).
28380 (* constant 5687 *)
28381 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)).
28382 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))).
28385 (* constant 5688 *)
28386 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))))))).
28387 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28390 (* constant 5689 *)
28391 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)))))).
28392 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))).
28395 (* constant 5690 *)
28396 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))))).
28397 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28400 (* constant 5691 *)
28401 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)))))).
28402 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28405 (* constant 5692 *)
28406 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)))))).
28407 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28410 (* constant 5693 *)
28411 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))))).
28412 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28415 (* constant 5694 *)
28416 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)))).
28417 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28420 (* constant 5695 *)
28421 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)))))))).
28422 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))).
28425 (* constant 5696 *)
28426 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)))))).
28427 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28430 (* constant 5697 *)
28431 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)))))).
28432 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28435 (* constant 5698 *)
28436 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))))))).
28437 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28440 (* constant 5699 *)
28441 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)))))).
28442 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28445 (* constant 5700 *)
28446 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)))))).
28447 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28450 (* constant 5701 *)
28451 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)))))).
28452 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28455 (* constant 5702 *)
28456 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))))))).
28457 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28460 (* constant 5703 *)
28461 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)))).
28462 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28465 (* constant 5704 *)
28466 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))).
28467 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))).
28470 (* constant 5705 *)
28471 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)))).
28472 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))).
28475 (* constant 5706 *)
28476 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)))).
28477 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))).
28480 (* constant 5707 *)
28481 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)))).
28482 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28485 (* constant 5708 *)
28486 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)))).
28487 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))).
28490 (* constant 5709 *)
28491 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))))).
28492 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28495 (* constant 5710 *)
28496 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))))))))).
28497 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28500 (* constant 5711 *)
28501 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))))))))).
28502 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28505 (* constant 5712 *)
28506 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))))))))).
28507 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28510 (* constant 5713 *)
28511 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))))))))).
28512 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28515 (* constant 5714 *)
28516 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))))))).
28517 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))).
28520 (* constant 5715 *)
28521 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)))))))).
28522 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))).
28525 (* constant 5716 *)
28526 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)))))))).
28527 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28530 (* constant 5717 *)
28531 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))))))).
28532 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28535 (* constant 5718 *)
28536 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))))))).
28537 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28540 (* constant 5719 *)
28541 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)))))))))).
28542 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28545 (* constant 5720 *)
28546 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)))))))))).
28547 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28550 (* constant 5721 *)
28551 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))))))))))).
28552 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))))).
28555 (* constant 5722 *)
28556 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))))))).
28557 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28560 (* constant 5723 *)
28561 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))))))).
28562 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28565 (* constant 5724 *)
28566 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)))))))))).
28567 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28570 (* constant 5725 *)
28571 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))))))))).
28572 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28575 (* constant 5726 *)
28576 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)))))))))).
28577 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28580 (* constant 5727 *)
28581 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)))))))))).
28582 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28585 (* constant 5728 *)
28586 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))))))))).
28587 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28590 (* constant 5729 *)
28591 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))))))))).
28592 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28595 (* constant 5730 *)
28596 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)))))))))).
28597 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28600 (* constant 5731 *)
28601 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))))))).
28602 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28605 (* constant 5732 *)
28606 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)))))))).
28607 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28610 (* constant 5733 *)
28611 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))))))))).
28612 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))))).
28615 (* constant 5734 *)
28616 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)))))))))).
28617 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28620 (* constant 5735 *)
28621 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)))))))).
28622 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28625 (* constant 5736 *)
28626 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)))))))).
28627 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28630 (* constant 5737 *)
28631 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))))))))).
28632 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))))).
28635 (* constant 5738 *)
28636 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))))))))).
28637 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))))).
28640 (* constant 5739 *)
28641 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))))))))).
28642 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28645 (* constant 5740 *)
28646 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))))))))).
28647 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28650 (* constant 5741 *)
28651 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))))))))).
28652 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28655 (* constant 5742 *)
28656 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)))))))).
28657 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28660 (* constant 5743 *)
28661 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)))))))).
28662 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28665 (* constant 5744 *)
28666 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))))))))).
28667 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28670 (* constant 5745 *)
28671 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)))))))))).
28672 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28675 (* constant 5746 *)
28676 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)))))))))).
28677 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28680 (* constant 5747 *)
28681 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)))))))).
28682 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28685 (* constant 5748 *)
28686 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))))))))).
28687 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28690 (* constant 5749 *)
28691 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))))))).
28692 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28695 (* constant 5750 *)
28696 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)))))))).
28697 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28700 (* constant 5751 *)
28701 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)))))))).
28702 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28705 (* constant 5752 *)
28706 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))))))))).
28707 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))))).
28710 (* constant 5753 *)
28711 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))))))))).
28712 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))))).
28715 (* constant 5754 *)
28716 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))))))))).
28717 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28720 (* constant 5755 *)
28721 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))))))))).
28722 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28725 (* constant 5756 *)
28726 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)))))))).
28727 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28730 (* constant 5757 *)
28731 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)))))))).
28732 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28735 (* constant 5758 *)
28736 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))))))))).
28737 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28740 (* constant 5759 *)
28741 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))))))))).
28742 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))))).
28745 (* constant 5760 *)
28746 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)))))))))).
28747 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28750 (* constant 5761 *)
28751 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)))))).
28752 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))))).
28755 (* constant 5762 *)
28756 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)))))).
28757 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))).
28760 (* constant 5763 *)
28761 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)))))).
28762 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))))).
28765 (* constant 5764 *)
28766 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)))).
28767 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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)))).
28770 (* constant 5765 *)
28771 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)))))))).
28772 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_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))))).
28775 (* constant 5766 *)
28776 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).
28777 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28780 (* constant 5767 *)
28781 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)).
28782 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)).
28785 (* constant 5768 *)
28786 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)).
28787 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)).
28790 (* constant 5769 *)
28791 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))))).
28792 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (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)))).
28795 (* constant 5770 *)
28796 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)))))).
28797 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (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)))))))).
28800 (* constant 5771 *)
28801 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)))).
28802 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc 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)))))).
28805 (* constant 5772 *)
28806 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)))))))).
28807 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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))))))).
28810 (* constant 5773 *)
28811 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)))))))).
28812 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))).
28815 (* constant 5774 *)
28816 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)))))))).
28817 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))).
28820 (* constant 5775 *)
28821 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)))))))).
28822 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))).
28825 (* constant 5776 *)
28826 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))))).
28827 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))).
28830 (* constant 5777 *)
28831 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))).
28832 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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))))).
28835 (* constant 5778 *)
28836 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)))).
28837 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc 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)))).
28840 (* constant 5779 *)
28841 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)))))))).
28842 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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))))))).
28845 (* constant 5780 *)
28846 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))))))))).
28847 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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))))))).
28850 (* constant 5781 *)
28851 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)))))))).
28852 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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))))))).
28855 (* constant 5782 *)
28856 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)))))))).
28857 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))))).
28860 (* constant 5783 *)
28861 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)))))))).
28862 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))))).
28865 (* constant 5784 *)
28866 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))))))))).
28867 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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))))))).
28870 (* constant 5785 *)
28871 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)))))))).
28872 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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))))))).
28875 (* constant 5786 *)
28876 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))))))))).
28877 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))))).
28880 (* constant 5787 *)
28881 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))))))))).
28882 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))))).
28885 (* constant 5788 *)
28886 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)))))))).
28887 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28890 (* constant 5789 *)
28891 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))))))).
28892 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28895 (* constant 5790 *)
28896 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))))))).
28897 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28900 (* constant 5791 *)
28901 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)))))))).
28902 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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))))))))).
28905 (* constant 5792 *)
28906 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))))))).
28907 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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))))))))).
28910 (* constant 5793 *)
28911 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)))))))))))).
28912 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28915 (* constant 5794 *)
28916 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))))))))).
28917 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28920 (* constant 5795 *)
28921 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)))))))))))).
28922 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28925 (* constant 5796 *)
28926 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)))))))))))).
28927 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))))).
28930 (* constant 5797 *)
28931 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)))))))).
28932 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28935 (* constant 5798 *)
28936 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)))))))))))).
28937 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28940 (* constant 5799 *)
28941 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)))))))).
28942 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28945 (* constant 5800 *)
28946 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))))))))).
28947 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28950 (* constant 5801 *)
28951 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)))))))))))).
28952 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28955 (* constant 5802 *)
28956 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))))))).
28957 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))))).
28960 (* constant 5803 *)
28961 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)))))).
28962 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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))))))).
28965 (* constant 5804 *)
28966 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)))))).
28967 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_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)))))).
28970 (* constant 5805 *)
28971 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)))).
28972 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc 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)))).
28975 (* constant 5806 *)
28976 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)))))))).
28977 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (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)))))).
28980 (* constant 5807 *)
28981 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))))))).
28982 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (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)))))))).
28985 (* constant 5808 *)
28986 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)))))).
28987 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (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))))))).
28990 (* constant 5809 *)
28991 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)))).
28992 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc 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))))))).
28995 (* constant 5810 *)
28996 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)))))).
28997 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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))))))).
29000 (* constant 5811 *)
29001 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))))))).
29002 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))).
29005 (* constant 5812 *)
29006 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))))))).
29007 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))).
29010 (* constant 5813 *)
29011 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)))))).
29012 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))).
29015 (* constant 5814 *)
29016 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))))).
29017 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))).
29020 (* constant 5815 *)
29021 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))).
29022 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_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)))))).
29025 (* constant 5816 *)
29026 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)))).
29027 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc 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)))).
29030 (* constant 5817 *)
29031 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))))).
29032 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc 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))))).
29035 (* constant 5818 *)
29036 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)))))))).
29037 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))).
29040 (* constant 5819 *)
29041 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))))))))))).
29042 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29045 (* constant 5820 *)
29046 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)))))))))).
29047 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29050 (* constant 5821 *)
29051 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)))))))))))))).
29052 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29055 (* constant 5822 *)
29056 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))))))))))))).
29057 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29060 (* constant 5823 *)
29061 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)))))))))))).
29062 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29065 (* constant 5824 *)
29066 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)))))))))))).
29067 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29070 (* constant 5825 *)
29071 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))))))))))).
29072 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29075 (* constant 5826 *)
29076 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))))))))))).
29077 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29080 (* constant 5827 *)
29081 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)))))))))).
29082 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29085 (* constant 5828 *)
29086 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)))))))))).
29087 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29090 (* constant 5829 *)
29091 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)))))))))).
29092 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29095 (* constant 5830 *)
29096 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))))))))))))).
29097 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29100 (* constant 5831 *)
29101 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))))))))))))).
29102 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29105 (* constant 5832 *)
29106 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)))))))))))).
29107 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29110 (* constant 5833 *)
29111 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))))))))))).
29112 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29115 (* constant 5834 *)
29116 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)))))))))).
29117 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29120 (* constant 5835 *)
29121 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))))))))))))).
29122 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29125 (* constant 5836 *)
29126 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))))))))))))).
29127 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29130 (* constant 5837 *)
29131 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)))))))))))).
29132 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29135 (* constant 5838 *)
29136 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)))))))))))).
29137 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29140 (* constant 5839 *)
29141 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))))))))))).
29142 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29145 (* constant 5840 *)
29146 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))))))))))).
29147 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29150 (* constant 5841 *)
29151 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)))))))))).
29152 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29155 (* constant 5842 *)
29156 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)))))))))).
29157 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29160 (* constant 5843 *)
29161 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)))))))))))).
29162 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29165 (* constant 5844 *)
29166 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)))))))))))).
29167 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29170 (* constant 5845 *)
29171 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))))))))))))).
29172 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29175 (* constant 5846 *)
29176 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)))))))))))).
29177 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29180 (* constant 5847 *)
29181 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)))))))))))).
29182 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29185 (* constant 5848 *)
29186 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)))))))))).
29187 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29190 (* constant 5849 *)
29191 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)))))))))).
29192 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29195 (* constant 5850 *)
29196 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)))))))))).
29197 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))).
29200 (* constant 5851 *)
29201 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)))))))))).
29202 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29205 (* constant 5852 *)
29206 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)))))))))).
29207 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29210 (* constant 5853 *)
29211 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)))))))))).
29212 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29215 (* constant 5854 *)
29216 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)))))))))))))).
29217 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29220 (* constant 5855 *)
29221 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)))))))))))).
29222 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29225 (* constant 5856 *)
29226 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)))))))))))).
29227 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29230 (* constant 5857 *)
29231 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)))))))))))).
29232 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29235 (* constant 5858 *)
29236 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))))))))))).
29237 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29240 (* constant 5859 *)
29241 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)))))))))).
29242 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29245 (* constant 5860 *)
29246 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))))))))))).
29247 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29250 (* constant 5861 *)
29251 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))))))))))).
29252 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29255 (* constant 5862 *)
29256 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)))))))))))).
29257 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))).
29260 (* constant 5863 *)
29261 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))))))))))))).
29262 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29265 (* constant 5864 *)
29266 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))))))))))))).
29267 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29270 (* constant 5865 *)
29271 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))))))))))))).
29272 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29275 (* constant 5866 *)
29276 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)))))))))).
29277 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29280 (* constant 5867 *)
29281 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))))))))).
29282 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29285 (* constant 5868 *)
29286 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)))))))).
29287 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_e_st_eq_landau_n_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)))))))).
29290 (* constant 5869 *)
29291 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))))))))))).
29292 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29295 (* constant 5870 *)
29296 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)))))))))).
29297 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29300 (* constant 5871 *)
29301 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))))))))))).
29302 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))).
29305 (* constant 5872 *)
29306 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)))))))))).
29307 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29310 (* constant 5873 *)
29311 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)))))))))))))).
29312 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29315 (* constant 5874 *)
29316 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))))))))))))).
29317 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29320 (* constant 5875 *)
29321 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))))))))))).
29322 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29325 (* constant 5876 *)
29326 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))))))))))))).
29327 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29330 (* constant 5877 *)
29331 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))))))))))).
29332 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29335 (* constant 5878 *)
29336 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))))))))))).
29337 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29340 (* constant 5879 *)
29341 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)))))))))).
29342 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29345 (* constant 5880 *)
29346 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)))))))))).
29347 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29350 (* constant 5881 *)
29351 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))))))))))).
29352 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29355 (* constant 5882 *)
29356 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)))))))))).
29357 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29360 (* constant 5883 *)
29361 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)))))))))).
29362 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29365 (* constant 5884 *)
29366 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)))))))))).
29367 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29370 (* constant 5885 *)
29371 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))))))))))))).
29372 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29375 (* constant 5886 *)
29376 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))))))))))))).
29377 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29380 (* constant 5887 *)
29381 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)))))))))))))).
29382 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29385 (* constant 5888 *)
29386 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)))))))))))))).
29387 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29390 (* constant 5889 *)
29391 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))))))))))))).
29392 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29395 (* constant 5890 *)
29396 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)))))))))))).
29397 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29400 (* constant 5891 *)
29401 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))))))))))).
29402 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29405 (* constant 5892 *)
29406 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))))))))))).
29407 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29410 (* constant 5893 *)
29411 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)))))))))).
29412 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29415 (* constant 5894 *)
29416 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))))))))))))).
29417 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29420 (* constant 5895 *)
29421 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))))))))))))).
29422 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29425 (* constant 5896 *)
29426 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))))))))))).
29427 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29430 (* constant 5897 *)
29431 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))))))))))).
29432 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29435 (* constant 5898 *)
29436 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))))))))))).
29437 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29440 (* constant 5899 *)
29441 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)))))))))).
29442 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29445 (* constant 5900 *)
29446 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)))))))))).
29447 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29450 (* constant 5901 *)
29451 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))))))))))).
29452 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29455 (* constant 5902 *)
29456 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)))))))))).
29457 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29460 (* constant 5903 *)
29461 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)))))))))).
29462 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29465 (* constant 5904 *)
29466 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)))))))))))).
29467 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29470 (* constant 5905 *)
29471 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))))))))))))).
29472 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29475 (* constant 5906 *)
29476 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))))))))))))).
29477 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29480 (* constant 5907 *)
29481 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))))))))))))).
29482 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29485 (* constant 5908 *)
29486 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)))))))))))).
29487 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29490 (* constant 5909 *)
29491 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)))))))))))).
29492 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29495 (* constant 5910 *)
29496 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)))))))))))).
29497 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29500 (* constant 5911 *)
29501 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)))))))))).
29502 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29505 (* constant 5912 *)
29506 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)))))))))).
29507 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29510 (* constant 5913 *)
29511 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)))))))))).
29512 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))).
29515 (* constant 5914 *)
29516 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)))))))))).
29517 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29520 (* constant 5915 *)
29521 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)))))))))).
29522 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29525 (* constant 5916 *)
29526 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)))))))))).
29527 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29530 (* constant 5917 *)
29531 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)))))))))).
29532 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29535 (* constant 5918 *)
29536 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)))))))))).
29537 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29540 (* constant 5919 *)
29541 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))))))))))))).
29542 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29545 (* constant 5920 *)
29546 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))))))))))))).
29547 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29550 (* constant 5921 *)
29551 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))))))))))).
29552 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29555 (* constant 5922 *)
29556 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)))))))))).
29557 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29560 (* constant 5923 *)
29561 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))))))))))).
29562 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29565 (* constant 5924 *)
29566 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))))))))))).
29567 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29570 (* constant 5925 *)
29571 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)))))))))).
29572 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29575 (* constant 5926 *)
29576 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)))))))))).
29577 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29580 (* constant 5927 *)
29581 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)))))))))).
29582 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29585 (* constant 5928 *)
29586 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))))))))))).
29587 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29590 (* constant 5929 *)
29591 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)))))))))).
29592 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29595 (* constant 5930 *)
29596 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))))))))))).
29597 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29600 (* constant 5931 *)
29601 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))))))))))).
29602 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29605 (* constant 5932 *)
29606 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)))))))))).
29607 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29610 (* constant 5933 *)
29611 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)))))))))).
29612 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29615 (* constant 5934 *)
29616 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)))))))))))).
29617 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29620 (* constant 5935 *)
29621 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)))))))))).
29622 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29625 (* constant 5936 *)
29626 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))))))))))).
29627 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29630 (* constant 5937 *)
29631 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)))))))))))).
29632 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29635 (* constant 5938 *)
29636 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)))))))))))).
29637 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29640 (* constant 5939 *)
29641 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)))))))))))).
29642 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29645 (* constant 5940 *)
29646 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))))))))))).
29647 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29650 (* constant 5941 *)
29651 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))))))))))).
29652 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29655 (* constant 5942 *)
29656 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)))))))))).
29657 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29660 (* constant 5943 *)
29661 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))))))))))).
29662 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29665 (* constant 5944 *)
29666 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))))))))).
29667 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29670 (* constant 5945 *)
29671 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))))))))))).
29672 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29675 (* constant 5946 *)
29676 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))))))))))).
29677 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29680 (* constant 5947 *)
29681 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)))))))))))).
29682 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29685 (* constant 5948 *)
29686 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))))))))))))).
29687 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29690 (* constant 5949 *)
29691 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)))))))))))))).
29692 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29695 (* constant 5950 *)
29696 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))))))))))))).
29697 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29700 (* constant 5951 *)
29701 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)))))))))))).
29702 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29705 (* constant 5952 *)
29706 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)))))))))))))).
29707 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29710 (* constant 5953 *)
29711 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)))))))))))))).
29712 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29715 (* constant 5954 *)
29716 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))))))))))))).
29717 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29720 (* constant 5955 *)
29721 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)))))))))))))).
29722 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29725 (* constant 5956 *)
29726 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))))))))))).
29727 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
29730 (* constant 5957 *)
29731 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))))))))))))).
29732 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29735 (* constant 5958 *)
29736 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)))))))))))).
29737 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29740 (* constant 5959 *)
29741 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)))))))))))))).
29742 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29745 (* constant 5960 *)
29746 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))))))))))))))).
29747 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29750 (* constant 5961 *)
29751 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))))))))))))))).
29752 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29755 (* constant 5962 *)
29756 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))))))))))))))).
29757 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29760 (* constant 5963 *)
29761 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))))))))))))))))).
29762 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))))))).
29765 (* constant 5964 *)
29766 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))))))))))))))))).
29767 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
29770 (* constant 5965 *)
29771 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)))))))))))))))).
29772 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))))))).
29775 (* constant 5966 *)
29776 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)))))))))))))))).
29777 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
29780 (* constant 5967 *)
29781 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))))))))))))))).
29782 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
29785 (* constant 5968 *)
29786 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)))))))))))))))).
29787 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
29790 (* constant 5969 *)
29791 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))))))))))))))).
29792 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29795 (* constant 5970 *)
29796 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)))))))))))))).
29797 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29800 (* constant 5971 *)
29801 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))))))))))))).
29802 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29805 (* constant 5972 *)
29806 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)))))))))))).
29807 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29810 (* constant 5973 *)
29811 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)))))))))).
29812 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29815 (* constant 5974 *)
29816 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))))))))))))).
29817 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29820 (* constant 5975 *)
29821 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))))))))))))).
29822 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
29825 (* constant 5976 *)
29826 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))))))))))))).
29827 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29830 (* constant 5977 *)
29831 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)))))))))))))).
29832 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29835 (* constant 5978 *)
29836 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))))))))))))).
29837 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29840 (* constant 5979 *)
29841 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)))))))))))).
29842 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29845 (* constant 5980 *)
29846 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))))))))))).
29847 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29850 (* constant 5981 *)
29851 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)))))))))))))).
29852 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29855 (* constant 5982 *)
29856 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))))))))))))).
29857 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29860 (* constant 5983 *)
29861 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))))))))))))))).
29862 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
29865 (* constant 5984 *)
29866 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)))))))))))))))).
29867 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
29870 (* constant 5985 *)
29871 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)))))))))))))))).
29872 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
29875 (* constant 5986 *)
29876 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)))))))))))))))).
29877 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
29880 (* constant 5987 *)
29881 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))))))))))))))).
29882 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
29885 (* constant 5988 *)
29886 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)))))))))))))))).
29887 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
29890 (* constant 5989 *)
29891 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)))))))))))))))).
29892 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
29895 (* constant 5990 *)
29896 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))))))))))))))))).
29897 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))))).
29900 (* constant 5991 *)
29901 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)))))))))))))))).
29902 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
29905 (* constant 5992 *)
29906 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))))))))))))))).
29907 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29910 (* constant 5993 *)
29911 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)))))))))))))).
29912 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29915 (* constant 5994 *)
29916 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)))))))))))))).
29917 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
29920 (* constant 5995 *)
29921 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)))))))))))))).
29922 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29925 (* constant 5996 *)
29926 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))))))))))))))).
29927 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29930 (* constant 5997 *)
29931 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)))))))))))))).
29932 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29935 (* constant 5998 *)
29936 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))))))))))))).
29937 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29940 (* constant 5999 *)
29941 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)))))))))))).
29942 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29945 (* constant 6000 *)
29946 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)))))))))))))).
29947 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29950 (* constant 6001 *)
29951 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)))))))))))))).
29952 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29955 (* constant 6002 *)
29956 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)))))))))))))).
29957 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29960 (* constant 6003 *)
29961 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))))))))))))))).
29962 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
29965 (* constant 6004 *)
29966 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)))))))))))))).
29967 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29970 (* constant 6005 *)
29971 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)))))))))))))).
29972 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29975 (* constant 6006 *)
29976 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))))))))))))))).
29977 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29980 (* constant 6007 *)
29981 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)))))))))))))))).
29982 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29985 (* constant 6008 *)
29986 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)))))))))))))))).
29987 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29990 (* constant 6009 *)
29991 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)))))))))))))).
29992 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
29995 (* constant 6010 *)
29996 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))))))))))))))).
29997 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30000 (* constant 6011 *)
30001 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)))))))))))))))).
30002 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30005 (* constant 6012 *)
30006 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))))))))))))).
30007 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30010 (* constant 6013 *)
30011 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))))))))))))))).
30012 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30015 (* constant 6014 *)
30016 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)))))))))))))))).
30017 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30020 (* constant 6015 *)
30021 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))))))))))))))).
30022 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30025 (* constant 6016 *)
30026 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)))))))))))))))).
30027 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30030 (* constant 6017 *)
30031 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))))))))))))))).
30032 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
30035 (* constant 6018 *)
30036 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))))))))))))))).
30037 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30040 (* constant 6019 *)
30041 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))))))))))))))).
30042 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30045 (* constant 6020 *)
30046 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))))))))))))))).
30047 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30050 (* constant 6021 *)
30051 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)))))))))))))))).
30052 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30055 (* constant 6022 *)
30056 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)))))))))))))))).
30057 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30060 (* constant 6023 *)
30061 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)))))))))))))))).
30062 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30065 (* constant 6024 *)
30066 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)))))))))))))))).
30067 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30070 (* constant 6025 *)
30071 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))))))))))))).
30072 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30075 (* constant 6026 *)
30076 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))))))))))))).
30077 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
30080 (* constant 6027 *)
30081 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))))))))))))).
30082 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
30085 (* constant 6028 *)
30086 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))))))))))))).
30087 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30090 (* constant 6029 *)
30091 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))))))))))))))).
30092 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30095 (* constant 6030 *)
30096 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)))))))))))))).
30097 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
30100 (* constant 6031 *)
30101 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)))))))))))))))).
30102 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
30105 (* constant 6032 *)
30106 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)))))))))))))))).
30107 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
30110 (* constant 6033 *)
30111 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))))))))))))))).
30112 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30115 (* constant 6034 *)
30116 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)))))))))))))).
30117 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30120 (* constant 6035 *)
30121 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)))))))))))))).
30122 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30125 (* constant 6036 *)
30126 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))))))))))))))))).
30127 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30130 (* constant 6037 *)
30131 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))))))))))))).
30132 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30135 (* constant 6038 *)
30136 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))))))))))))).
30137 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30140 (* constant 6039 *)
30141 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)))))))))))))).
30142 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30145 (* constant 6040 *)
30146 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))))))))))))))).
30147 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30150 (* constant 6041 *)
30151 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)))))))))))))).
30152 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30155 (* constant 6042 *)
30156 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))))))))))))))).
30157 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30160 (* constant 6043 *)
30161 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)))))))))))))).
30162 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30165 (* constant 6044 *)
30166 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)))))))))))))))).
30167 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30170 (* constant 6045 *)
30171 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))))))))))))))).
30172 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30175 (* constant 6046 *)
30176 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))))))))))))))).
30177 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30180 (* constant 6047 *)
30181 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))))))))))))))).
30182 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30185 (* constant 6048 *)
30186 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))))))))))))))).
30187 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30190 (* constant 6049 *)
30191 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))))))))))))).
30192 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30195 (* constant 6050 *)
30196 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))))))))))))))).
30197 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30200 (* constant 6051 *)
30201 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))))))))))))))).
30202 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30205 (* constant 6052 *)
30206 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)))))))))))))).
30207 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30210 (* constant 6053 *)
30211 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)))))))))))))).
30212 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30215 (* constant 6054 *)
30216 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))))))))))))))).
30217 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30220 (* constant 6055 *)
30221 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)))))))))))))))).
30222 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30225 (* constant 6056 *)
30226 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)))))))))))))))).
30227 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30230 (* constant 6057 *)
30231 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)))))))))))))))).
30232 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30235 (* constant 6058 *)
30236 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)))))))))))))))).
30237 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30240 (* constant 6059 *)
30241 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)))))))))))))))).
30242 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30245 (* constant 6060 *)
30246 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)))))))))))))))).
30247 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30250 (* constant 6061 *)
30251 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))))))))))))))))).
30252 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30255 (* constant 6062 *)
30256 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)))))))))))))))))).
30257 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
30260 (* constant 6063 *)
30261 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))))))))))))))))).
30262 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))))).
30265 (* constant 6064 *)
30266 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)))))))))))))).
30267 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30270 (* constant 6065 *)
30271 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)))))))))))))).
30272 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30275 (* constant 6066 *)
30276 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)))))))))))))).
30277 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30280 (* constant 6067 *)
30281 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)))))))))))))).
30282 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30285 (* constant 6068 *)
30286 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))))))))))))))))).
30287 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30290 (* constant 6069 *)
30291 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))))))))))))))).
30292 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30295 (* constant 6070 *)
30296 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)))))))))))))))).
30297 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30300 (* constant 6071 *)
30301 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))))))))))))))).
30302 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30305 (* constant 6072 *)
30306 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))))))))))))))).
30307 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30310 (* constant 6073 *)
30311 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)))))))))))))))).
30312 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30315 (* constant 6074 *)
30316 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)))))))))))))))).
30317 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
30320 (* constant 6075 *)
30321 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)))))))))))))))).
30322 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30325 (* constant 6076 *)
30326 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)))))))))))))))).
30327 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30330 (* constant 6077 *)
30331 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)))))))))))))))).
30332 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30335 (* constant 6078 *)
30336 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)))))))))))))))).
30337 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30340 (* constant 6079 *)
30341 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)))))))))))))))).
30342 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30345 (* constant 6080 *)
30346 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)))))))))))))).
30347 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30350 (* constant 6081 *)
30351 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))))))))))))))).
30352 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30355 (* constant 6082 *)
30356 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)))))))))))))))).
30357 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30360 (* constant 6083 *)
30361 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)))))))))))))))).
30362 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
30365 (* constant 6084 *)
30366 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))))))))))))))).
30367 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
30370 (* constant 6085 *)
30371 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))))))))))))))))).
30372 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30375 (* constant 6086 *)
30376 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))))))))))))))).
30377 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30380 (* constant 6087 *)
30381 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))))))))))))))))).
30382 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))))).
30385 (* constant 6088 *)
30386 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))))))))))))))))).
30387 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))))).
30390 (* constant 6089 *)
30391 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))))))))))))))).
30392 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
30395 (* constant 6090 *)
30396 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))))))))))))))).
30397 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
30400 (* constant 6091 *)
30401 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))))))))))))))).
30402 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
30405 (* constant 6092 *)
30406 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))))))))))))))))).
30407 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30410 (* constant 6093 *)
30411 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)))))))))))))))).
30412 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30415 (* constant 6094 *)
30416 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))))))))))))))).
30417 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30420 (* constant 6095 *)
30421 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)))))))))))))).
30422 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))))).
30425 (* constant 6096 *)
30426 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))))))))))))))).
30427 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30430 (* constant 6097 *)
30431 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)))))))))))))))).
30432 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))))).
30435 (* constant 6098 *)
30436 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)))))))))))))))).
30437 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30440 (* constant 6099 *)
30441 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)))))))))))))))).
30442 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30445 (* constant 6100 *)
30446 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))))))))))))))).
30447 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30450 (* constant 6101 *)
30451 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))))))))))))))))).
30452 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30455 (* constant 6102 *)
30456 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))))))))))))))).
30457 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30460 (* constant 6103 *)
30461 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)))))))))))))).
30462 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30465 (* constant 6104 *)
30466 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))))))))))))))))).
30467 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30470 (* constant 6105 *)
30471 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))))))))))))).
30472 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))))).
30475 (* constant 6106 *)
30476 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)))))))))))).
30477 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))))).
30480 (* constant 6107 *)
30481 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))))))))))).
30482 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))))).
30485 (* constant 6108 *)
30486 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)))))))))).
30487 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))))).
30490 (* constant 6109 *)
30491 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))))))))).
30492 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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)))))))))).
30495 (* constant 6110 *)
30496 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)))))))).
30497 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_st_eq_landau_n_nat) => (fun (p:l_e_st_eq_landau_n_rt_rp_r_c_8283_prop2 q a c x) => (fun (s:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x))) => (fun (f:(forall (t:l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)), l_e_st_eq_landau_n_rt_rp_r_c_cx)) => (fun (b:l_e_bijective (l_e_st_eq_landau_n_1to (l_e_st_eq_landau_n_rt_rp_r_c_8283_xp1 q a c x)) (l_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))))))))).
30500 (* constant 6111 *)
30501 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)))))).
30502 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_e_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)))))))).
30505 (* constant 6112 *)
30506 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)))))).
30507 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (fun (c:l_e_st_eq_landau_n_rt_rp_r_c_commut q) => (fun (x:l_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)))))).
30510 (* constant 6113 *)
30511 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)))).
30512 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc 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)))).
30515 (* constant 6114 *)
30516 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)))))))).
30517 exact (fun (q:(forall (t:l_e_st_eq_landau_n_rt_rp_r_c_cx), (forall (u:l_e_st_eq_landau_n_rt_rp_r_c_cx), l_e_st_eq_landau_n_rt_rp_r_c_cx))) => (fun (a:l_e_st_eq_landau_n_rt_rp_r_c_assoc q) => (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))))))).
30520 (* constant 6115 *)
30521 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))))).
30522 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))))).
30525 (* constant 6116 *)
30526 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)))))).
30527 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)))))).
30530 (* constant 6117 *)
30531 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))))))).
30532 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)))))).
30535 (* constant 6118 *)
30536 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)))))).
30537 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)))))).
30540 (* constant 6119 *)
30541 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))))))).
30542 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)))))).
30545 (* constant 6120 *)
30546 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)))))))).
30547 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)))))))).
30550 (* constant 6121 *)
30551 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)))))))).
30552 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))))))).
30555 (* constant 6122 *)
30556 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))))))))).
30557 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))))))).
30560 (* constant 6123 *)
30561 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))))))).
30562 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))))))).
30565 (* constant 6124 *)
30566 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))))))).
30567 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)))))).
30570 (* constant 6125 *)
30571 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))))))).
30572 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)))))))).
30575 (* constant 6126 *)
30576 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))))))))))))))))).
30577 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)))))))))))))))))).
30580 (* constant 6127 *)
30581 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))))))))))))))))).
30582 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)))))))))))))))))).
30585 (* constant 6128 *)
30586 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)).
30587 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)).
30590 (* constant 6129 *)
30591 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)).
30592 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)).
30595 (* constant 6130 *)
30596 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).
30597 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).
30600 (* constant 6131 *)
30601 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)))))))))))))))))).
30602 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)))))))))))))))))))).
30605 (* constant 6132 *)
30606 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)))))))))))))))))).
30607 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)))))))))))))))))).
30610 (* constant 6133 *)
30611 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))))))))))))))))).
30612 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)))))))))))))))))).
30615 (* constant 6134 *)
30616 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))))))))))))))))).
30617 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))))))))))))))))).
30620 (* constant 6135 *)
30621 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)))))))))))))).
30622 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))))))))))))).
30625 (* constant 6136 *)
30626 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))))))))))))).
30627 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))))))))))))).
30630 (* constant 6137 *)
30631 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)))))))))))))).
30632 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)))))))))))))))).
30635 (* constant 6138 *)
30636 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))))))))))))).
30637 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))))))))))))).
30640 (* constant 6139 *)
30641 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))))))))))))).
30642 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))))))))))))).
30645 (* constant 6140 *)
30646 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))))))))))))).
30647 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)))))))))))))).
30650 (* constant 6141 *)
30651 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))))))))))))))).
30652 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))))))))))))))).
30655 (* constant 6142 *)
30656 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)))))))))))))).
30657 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)))))))))))))).
30660 (* constant 6143 *)
30661 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)))))))))))))).
30662 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)))))))))))))).
30665 (* constant 6144 *)
30666 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)))))))))))))).
30667 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)))))))))))))).
30670 (* constant 6145 *)
30671 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)))))))))))))).
30672 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)))))))))))))).
30675 (* constant 6146 *)
30676 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)))))))))))))).
30677 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)))))))))))))).
30680 (* constant 6147 *)
30681 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))))))))))))))).
30682 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))))))))))))))).
30685 (* constant 6148 *)
30686 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)))))))))))))).
30687 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)))))))))))))).
30690 (* constant 6149 *)
30691 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)))))))))))))))).
30692 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)))))))))))))).
30695 (* constant 6150 *)
30696 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)))))))))))))))).
30697 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)))))))))))))))).
30700 (* constant 6151 *)
30701 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))))))))))))))))).
30702 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)))))))))))))))).
30705 (* constant 6152 *)
30706 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))))))))))))))).
30707 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))))))))))))))).
30710 (* constant 6153 *)
30711 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)))))))))))))))).
30712 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))))))))))))))).
30715 (* constant 6154 *)
30716 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))))))))))))))))).
30717 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))))))))))))))).
30720 (* constant 6155 *)
30721 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))))))))))))))))).
30722 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))))))))))))))).
30725 (* constant 6156 *)
30726 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))))))))))))))))).
30727 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)))))))))))))))).
30730 (* constant 6157 *)
30731 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))))))))))))))))).
30732 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))))))))))))))).
30735 (* constant 6158 *)
30736 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))))))))))))))).
30737 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))))))))))))))).
30740 (* constant 6159 *)
30741 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)))))))))))))))).
30742 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)))))))))))))))).
30745 (* constant 6160 *)
30746 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)))))))))))))))).
30747 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))))))))))))))).
30750 (* constant 6161 *)
30751 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))))))))))))))).
30752 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))))))))))))))).
30755 (* constant 6162 *)
30756 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)))))))))))))))).
30757 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))))))))))))))).
30760 (* constant 6163 *)
30761 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)))))))))))))))).
30762 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))))))))))))))).
30765 (* constant 6164 *)
30766 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)))))))))))))).
30767 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))))))))))))))).
30770 (* constant 6165 *)
30771 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))))))))))))))).
30772 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)))))))))))))).
30775 (* constant 6166 *)
30776 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))))))))))))))).
30777 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)))))))))))))).
30780 (* constant 6167 *)
30781 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)))))))))))))).
30782 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)))))))))))))).
30785 (* constant 6168 *)
30786 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))))))))))))))).
30787 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))))))))))))))).
30790 (* constant 6169 *)
30791 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))))))))))))))).
30792 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))))))))))))))).
30795 (* constant 6170 *)
30796 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))))))))))))))).
30797 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)))))))))))))).
30800 (* constant 6171 *)
30801 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)))))))))))))).
30802 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)))))))))))))).
30805 (* constant 6172 *)
30806 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))))))))))))))).
30807 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)))))))))))))).
30810 (* constant 6173 *)
30811 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)))))))))))))))).
30812 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)))))))))))))))).
30815 (* constant 6174 *)
30816 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))))))))))))))))).
30817 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)))))))))))))))).
30820 (* constant 6175 *)
30821 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))))))))))))))).
30822 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))))))))))))))).
30825 (* constant 6176 *)
30826 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)))))))))))))))).
30827 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))))))))))))))).
30830 (* constant 6177 *)
30831 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))))))))))))))))).
30832 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))))))))))))))).
30835 (* constant 6178 *)
30836 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))))))))))))))).
30837 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))))))))))))))).
30840 (* constant 6179 *)
30841 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)))))))))))))))).
30842 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))))))))))))))).
30845 (* constant 6180 *)
30846 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))))))))))))))).
30847 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)))))))))))))).
30850 (* constant 6181 *)
30851 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))))))))))))))).
30852 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)))))))))))))).
30855 (* constant 6182 *)
30856 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)))))))))))))).
30857 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)))))))))))))).
30860 (* constant 6183 *)
30861 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)))))))))))))).
30862 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))))))))))))))).
30865 (* constant 6184 *)
30866 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)))))))))))))))).
30867 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))))))))))))))).
30870 (* constant 6185 *)
30871 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)))))))))))))))).
30872 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))))))))))))))).
30875 (* constant 6186 *)
30876 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)))))))))))))).
30877 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))))))))))))))).
30880 (* constant 6187 *)
30881 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))))))))))))))).
30882 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))))))))))))))).
30885 (* constant 6188 *)
30886 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))))))))))))))).
30887 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))))))))))))))).
30890 (* constant 6189 *)
30891 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))))))))))))))).
30892 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)))))))))))))).
30895 (* constant 6190 *)
30896 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)))))))))))))).
30897 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)))))))))))))).
30900 (* constant 6191 *)
30901 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))))))))))))))).
30902 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)))))))))))))).
30905 (* constant 6192 *)
30906 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))))))))))))))).
30907 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)))))))))))))).
30910 (* constant 6193 *)
30911 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))))))))))))))).
30912 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)))))))))))))).
30915 (* constant 6194 *)
30916 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)).
30917 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)).
30920 (* constant 6195 *)
30921 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)).
30922 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)).
30925 (* constant 6196 *)
30926 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).
30927 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).
30930 (* constant 6197 *)
30931 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))))))))))))))).
30932 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))))))))))))))).
30935 (* constant 6198 *)
30936 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))))))))))))))).
30937 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))))))))))))))).
30940 (* constant 6199 *)
30941 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))))))))))))))).
30942 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))))))))))))))).
30945 (* constant 6200 *)
30946 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)))))))))))))).
30947 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))))))))))))))).
30950 (* constant 6201 *)
30951 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)))))))))))))).
30952 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))))))))))))))).
30955 (* constant 6202 *)
30956 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)))))))))))).
30957 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))))))))))).
30960 (* constant 6203 *)
30961 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))))))))))).
30962 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))))))))))).
30965 (* constant 6204 *)
30966 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))))))))))).
30967 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)))))))))))).
30970 (* constant 6205 *)
30971 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))))))))))).
30972 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))))))))))).
30975 (* constant 6206 *)
30976 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)))))))))).
30977 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)))))))))).
30980 (* constant 6207 *)
30981 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)))))))))).
30982 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))))))))))).
30985 (* constant 6208 *)
30986 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))))))))))).
30987 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))))))))))).
30990 (* constant 6209 *)
30991 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))))))))))).
30992 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))))))))))).
30995 (* constant 6210 *)
30996 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))))))))))).
30997 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))))))))))).
31000 (* constant 6211 *)
31001 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))))))))))).
31002 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))))))))))).
31005 (* constant 6212 *)
31006 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))))))))))))).
31007 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))))))))))))).
31010 (* constant 6213 *)
31011 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)))))))))))).
31012 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))))))))))).
31015 (* constant 6214 *)
31016 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)))))))))))))).
31017 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)))))))))))).
31020 (* constant 6215 *)
31021 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)))))))))))).
31022 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))))))))))))).
31025 (* constant 6216 *)
31026 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))))))))))).
31027 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)))))))))))).
31030 (* constant 6217 *)
31031 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))))))))))).
31032 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))))))))))).
31035 (* constant 6218 *)
31036 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)))))))))))).
31037 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))))))))))))).
31040 (* constant 6219 *)
31041 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)))))))))))).
31042 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)))))))))))).
31045 (* constant 6220 *)
31046 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))))))))))).
31047 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)))))))))))).
31050 (* constant 6221 *)
31051 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)))))))))))).
31052 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)))))))))))).
31055 (* constant 6222 *)
31056 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)))))))))))).
31057 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))))))))))).
31060 (* constant 6223 *)
31061 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)))))))))))).
31062 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))))))))))).
31065 (* constant 6224 *)
31066 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)))))))))))).
31067 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))))))))))).
31070 (* constant 6225 *)
31071 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)))))))))))).
31072 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))))))))))).
31075 (* constant 6226 *)
31076 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)))))))))))).
31077 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)))))))))))).
31080 (* constant 6227 *)
31081 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))))))))))).
31082 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)))))))))))).
31085 (* constant 6228 *)
31086 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)))))))))))).
31087 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)))))))))))).
31090 (* constant 6229 *)
31091 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))))))))))).
31092 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))))))))))).
31095 (* constant 6230 *)
31096 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)))))))))))).
31097 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))))))))))).
31100 (* constant 6231 *)
31101 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))))))))))).
31102 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)))))))))).
31105 (* constant 6232 *)
31106 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))))))))))).
31107 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))))))))))).
31110 (* constant 6233 *)
31111 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)))))))))))).
31112 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)))))))))))).
31115 (* constant 6234 *)
31116 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))))))))))))).
31117 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)))))))))))).
31120 (* constant 6235 *)
31121 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))))))))))))).
31122 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))))))))))))).
31125 (* constant 6236 *)
31126 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))))))))))))).
31127 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))))))))))))).
31130 (* constant 6237 *)
31131 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)))))))))))).
31132 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))))))))))))).
31135 (* constant 6238 *)
31136 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)))))))))))).
31137 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))))))))))))).
31140 (* constant 6239 *)
31141 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)))))))))))).
31142 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)))))))))))).
31145 (* constant 6240 *)
31146 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)))))))))))))))).
31147 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)))))))))))))).
31150 (* constant 6241 *)
31151 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))))))))))))))).
31152 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))))))))))))))).
31155 (* constant 6242 *)
31156 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))))))))))))))).
31157 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)))))))))))))))).
31160 (* constant 6243 *)
31161 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))))))))))))))))).
31162 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)))))))))))))))).
31165 (* constant 6244 *)
31166 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)))))))))))))))).
31167 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)))))))))))))))).
31170 (* constant 6245 *)
31171 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)))))))))))))))).
31172 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)))))))))))))))).
31175 (* constant 6246 *)
31176 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))))))))))))))).
31177 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)))))))))))))))).
31180 (* constant 6247 *)
31181 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))))))))))))))))).
31182 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)))))))))))))))).
31185 (* constant 6248 *)
31186 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))))))))))))))))).
31187 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)))))))))))))))).
31190 (* constant 6249 *)
31191 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))))))))))))))).
31192 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)))))))))))))))).
31195 (* constant 6250 *)
31196 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))))))))))))))))).
31197 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)))))))))))))))).
31200 (* constant 6251 *)
31201 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)))))))))))))))).
31202 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))))))))))))))).
31205 (* constant 6252 *)
31206 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))))))))))))))))).
31207 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))))))))))))))).
31210 (* constant 6253 *)
31211 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))))))))))))))).
31212 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))))))))))))))).
31215 (* constant 6254 *)
31216 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))).
31217 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))).
31220 (* constant 6255 *)
31221 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))).
31222 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))).
31225 (* constant 6256 *)
31226 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))).
31227 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)))).
31230 (* constant 6257 *)
31231 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))).
31232 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)))).
31235 (* constant 6258 *)
31236 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)).
31237 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))).
31240 (* constant 6259 *)
31241 Definition l_e_st_eq_landau_n_rt_rp_r_c_8287_prop5 : (forall (x:l_e_st_eq_landau_n_nat), Prop).
31242 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)).
31245 (* constant 6260 *)
31246 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))).
31247 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).
31250 (* constant 6261 *)
31251 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)))).
31252 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)).
31255 (* constant 6262 *)
31256 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)))).
31257 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)).
31260 (* constant 6263 *)
31261 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)))).
31262 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)).
31265 (* constant 6264 *)
31266 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)))).
31267 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)).
31270 (* constant 6265 *)
31271 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).
31272 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)).
31275 (* constant 6266 *)
31276 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.
31277 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).
31280 (* constant 6267 *)
31281 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)))).
31282 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)))).
31285 (* constant 6268 *)
31286 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)))).
31287 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))).
31290 (* constant 6269 *)
31291 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))))).
31292 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))))).
31295 (* constant 6270 *)
31296 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))))).
31297 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))))).
31300 (* constant 6271 *)
31301 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))))))))).
31302 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))))).
31305 (* constant 6272 *)
31306 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))))))).
31307 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))))))).
31310 (* constant 6273 *)
31311 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))))).
31312 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)))))))).
31315 (* constant 6274 *)
31316 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))))))).
31317 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))))))))).
31320 (* constant 6275 *)
31321 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))))))).
31322 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)))))).
31325 (* constant 6276 *)
31326 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))))))).
31327 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)))))).
31330 (* constant 6277 *)
31331 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)))))).
31332 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)))))).
31335 (* constant 6278 *)
31336 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))))))).
31337 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)))))).
31340 (* constant 6279 *)
31341 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))))))).
31342 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)))))).
31345 (* constant 6280 *)
31346 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))))))).
31347 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))))).
31350 (* constant 6281 *)
31351 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)))))).
31352 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))))))).
31355 (* constant 6282 *)
31356 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))))))).
31357 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)))))).
31360 (* constant 6283 *)
31361 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))))))).
31362 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)))))).
31365 (* constant 6284 *)
31366 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))))).
31367 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)))))).
31370 (* constant 6285 *)
31371 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))).
31372 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))))).
31375 (* constant 6286 *)
31376 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))).
31377 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))).
31380 (* constant 6287 *)
31381 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))).
31382 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))).
31385 (* constant 6288 *)
31386 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))))).
31387 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)).
31390 (* constant 6289 *)
31391 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)).
31392 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)))).
31395 (* constant 6290 *)
31396 Definition l_e_st_eq_landau_n_rt_rp_r_c_8288_prop2 : (forall (x:l_e_st_eq_landau_n_nat), Prop).
31397 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)).
31400 (* constant 6291 *)
31401 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))).
31402 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).
31405 (* constant 6292 *)
31406 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)))).
31407 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)).
31410 (* constant 6293 *)
31411 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)).
31412 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)).
31415 (* constant 6294 *)
31416 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)).
31417 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)).
31420 (* constant 6295 *)
31421 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).
31422 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)).
31425 (* constant 6296 *)
31426 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.
31427 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).
31430 (* constant 6297 *)
31431 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)))).
31432 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)))).
31435 (* constant 6298 *)
31436 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)))).
31437 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))).
31440 (* constant 6299 *)
31441 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))))))).
31442 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))).
31445 (* constant 6300 *)
31446 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))).
31447 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)))))).
31450 (* constant 6301 *)
31451 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)))))))).
31452 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)))).
31455 (* constant 6302 *)
31456 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))))).
31457 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)))))).
31460 (* constant 6303 *)
31461 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))))).
31462 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)))).
31465 (* constant 6304 *)
31466 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)))).
31467 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)))).
31470 (* constant 6305 *)
31471 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)))).
31472 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)))).
31475 (* constant 6306 *)
31476 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))))).
31477 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)))).
31480 (* constant 6307 *)
31481 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)))).
31482 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))))).
31485 (* constant 6308 *)
31486 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))))).
31487 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)))).
31490 (* constant 6309 *)
31491 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))))))).
31492 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))).
31495 (* constant 6310 *)
31496 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))))).
31497 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)))))).
31500 (* constant 6311 *)
31501 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))).
31502 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))))).
31505 (* constant 6312 *)
31506 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)))).
31507 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)))).
31510 (* constant 6313 *)
31511 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)))).
31512 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)))).
31515 (* constant 6314 *)
31516 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))).
31517 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)))).
31520 (* constant 6315 *)
31521 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))).
31522 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))).
31525 (* constant 6316 *)
31526 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))).
31527 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))).
31530 (* constant 6317 *)
31531 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)))).
31532 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)).
31535 (* constant 6318 *)
31536 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)).
31537 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)).
31540 (* constant 6319 *)
31541 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)).
31542 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))).
31545 (* constant 6320 *)
31546 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)).
31547 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))).
31550 (* constant 6321 *)
31551 Definition l_e_st_eq_landau_n_rt_rp_r_c_8289_prop4 : (forall (x:l_e_st_eq_landau_n_nat), Prop).
31552 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)).
31555 (* constant 6322 *)
31556 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))).
31557 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).
31560 (* constant 6323 *)
31561 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)).
31562 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)).
31565 (* constant 6324 *)
31566 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)).
31567 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))).
31570 (* constant 6325 *)
31571 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))))).
31572 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)))).
31575 (* constant 6326 *)
31576 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)))).
31577 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)))).
31580 (* constant 6327 *)
31581 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)).
31582 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)))).
31585 (* constant 6328 *)
31586 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).
31587 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)).
31590 (* constant 6329 *)
31591 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.
31592 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).
31595 (* constant 6330 *)
31596 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)))).
31597 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)))).
31600 (* constant 6331 *)
31601 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)))).
31602 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))).
31605 (* constant 6332 *)
31606 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))))))).
31607 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))).
31610 (* constant 6333 *)
31611 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)))).
31612 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)))).
31615 (* constant 6334 *)
31616 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))))).
31617 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))))).
31620 (* constant 6335 *)
31621 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)))))).
31622 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))))).
31625 (* constant 6336 *)
31626 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))))))).
31627 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))))))).
31630 (* constant 6337 *)
31631 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))))).
31632 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))))))).
31635 (* constant 6338 *)
31636 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))))).
31637 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))))).
31640 (* constant 6339 *)
31641 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)))).
31642 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))))).
31645 (* constant 6340 *)
31646 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))))))).
31647 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))))))).
31650 (* constant 6341 *)
31651 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))))))).
31652 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)))))))).
31655 (* constant 6342 *)
31656 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))))))).
31657 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))))))).
31660 (* constant 6343 *)
31661 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)))))))))).
31662 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)))))))).
31665 (* constant 6344 *)
31666 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)))))))))).
31667 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))))))))).
31670 (* constant 6345 *)
31671 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)))))))).
31672 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)))))))).
31675 (* constant 6346 *)
31676 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)))))))).
31677 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)))))))).
31680 (* constant 6347 *)
31681 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))))))).
31682 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)))))))).
31685 (* constant 6348 *)
31686 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))))))).
31687 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)))))))).
31690 (* constant 6349 *)
31691 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))))))))).
31692 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)))))))).
31695 (* constant 6350 *)
31696 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))))))))).
31697 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)))))))).
31700 (* constant 6351 *)
31701 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))))))))).
31702 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)))))))).
31705 (* constant 6352 *)
31706 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))))))))).
31707 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)))))))).
31710 (* constant 6353 *)
31711 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))))))).
31712 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))))))).
31715 (* constant 6354 *)
31716 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)))))))).
31717 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)))))))).
31720 (* constant 6355 *)
31721 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)))))))).
31722 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)))))))).
31725 (* constant 6356 *)
31726 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))))))).
31727 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)))))))).
31730 (* constant 6357 *)
31731 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)))))).
31732 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))))))).
31735 (* constant 6358 *)
31736 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)))))).
31737 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))))))).
31740 (* constant 6359 *)
31741 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)))).
31742 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)))))).
31745 (* constant 6360 *)
31746 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))).
31747 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)))).
31750 (* constant 6361 *)
31751 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))).
31752 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))).
31755 (* constant 6362 *)
31756 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))).
31757 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))).
31760 (* constant 6363 *)
31761 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)))).
31762 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)).
31765 (* constant 6364 *)
31766 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)))).
31767 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))).
31770 (* constant 6365 *)
31771 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)))).
31772 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)))).
31775 (* constant 6366 *)
31776 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)))).
31777 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))))).
31780 (* constant 6367 *)
31781 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))))).
31782 exact (fun (x: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))))).
31785 (* constant 6368 *)
31786 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))))).
31787 exact (fun (x: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)))))).
31790 (* constant 6369 *)
31791 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))))).
31792 exact (fun (x: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)))))).
31795 (* constant 6370 *)
31796 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)))))))))))).
31797 exact (fun (x:l_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)))))))))))).
31800 (* constant 6371 *)
31801 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)))))))))))).
31802 exact (fun (x:l_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)))))))))))).
31805 (* constant 6372 *)
31806 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))))))))))))).
31807 exact (fun (x:l_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)))))))))))).
31810 (* constant 6373 *)
31811 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))))))))))))).
31812 exact (fun (x:l_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))))))))))))).
31815 (* constant 6374 *)
31816 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))))))))))))).
31817 exact (fun (x:l_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))))))))))))).
31820 (* constant 6375 *)
31821 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)))))))))))))).
31822 exact (fun (x:l_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))))))))))))).
31825 (* constant 6376 *)
31826 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))))))))))))))).
31827 exact (fun (x:l_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))))))))))))).
31830 (* constant 6377 *)
31831 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))))))))))))).
31832 exact (fun (x:l_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))))))))))))).
31835 (* constant 6378 *)
31836 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))))))).
31837 exact (fun (x: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)))))).
31840 (* constant 6379 *)
31841 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)))))))).
31842 exact (fun (x: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))))))).
31845 (* constant 6380 *)
31846 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)))))).
31847 exact (fun (x: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))))))).
31850 (* constant 6381 *)
31851 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)))).
31852 exact (fun (x: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))).
31855 (* constant 6382 *)
31856 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)))))).
31857 exact (fun (x: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))))).
31860 (* constant 6383 *)
31861 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))))))).
31862 exact (fun (x: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)))))).
31865 (* constant 6384 *)
31866 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))))).
31867 exact (fun (x: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)))))).
31870 (* constant 6385 *)
31871 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))))).
31872 exact (fun (x: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)))))).
31875 (* constant 6386 *)
31876 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))))).
31877 exact (fun (x: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)))))).
31880 (* constant 6387 *)
31881 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)))))))))))).
31882 exact (fun (x:l_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))))))))))))).
31885 (* constant 6388 *)
31886 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)))))))))))).
31887 exact (fun (x:l_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))))))))))))).
31890 (* constant 6389 *)
31891 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))))))))))))).
31892 exact (fun (x:l_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))))))))))))).
31895 (* constant 6390 *)
31896 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))))))))))))).
31897 exact (fun (x:l_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))))))))))))).
31900 (* constant 6391 *)
31901 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))))))).
31902 exact (fun (x: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)))))).
31905 (* constant 6392 *)
31906 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)))).
31907 exact (fun (x: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)))))).
31910 (* constant 6393 *)
31911 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)))))).
31912 exact (fun (x: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))))).
31915 (* constant 6394 *)
31916 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))))).
31917 exact (fun (x: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))))).
31920 (* constant 6395 *)
31921 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))))))))))).
31922 exact (fun (x:l_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))))))))))).
31925 (* constant 6396 *)
31926 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)))))))))))).
31927 exact (fun (x:l_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))))))))))).
31930 (* constant 6397 *)
31931 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)))))))))))).
31932 exact (fun (x:l_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))))))))))))).
31935 (* constant 6398 *)
31936 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)))))))))))).
31937 exact (fun (x:l_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)))))))))))).
31940 (* constant 6399 *)
31941 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)))))))))))).
31942 exact (fun (x:l_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))))))))))).
31945 (* constant 6400 *)
31946 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))))))))))).
31947 exact (fun (x:l_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))))))))))).
31950 (* constant 6401 *)
31951 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))))))))))).
31952 exact (fun (x:l_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)))))))))))).
31955 (* constant 6402 *)
31956 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)))))))))))).
31957 exact (fun (x:l_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)))))))))))).
31960 (* constant 6403 *)
31961 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))))))))))).
31962 exact (fun (x:l_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))))))))))).
31965 (* constant 6404 *)
31966 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)))).
31967 exact (fun (x: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))))))).
31970 (* constant 6405 *)
31971 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)))))).
31972 exact (fun (x: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))))).
31975 (* constant 6406 *)
31976 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)))))).
31977 exact (fun (x: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))))).
31980 (* constant 6407 *)
31981 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))))).
31982 exact (fun (x: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))))))).
31985 (* constant 6408 *)
31986 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)))))).
31987 exact (fun (x: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)))))).
31990 (* constant 6409 *)
31991 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))))))).
31992 exact (fun (x: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))))).
31995 (* constant 6410 *)
31996 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)))))).
31997 exact (fun (x: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))))))).
32000 (* constant 6411 *)
32001 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))))).
32002 exact (fun (x: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))))).
32005 (* constant 6412 *)
32006 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))))).
32007 exact (fun (x: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))))).
32010 (* constant 6413 *)
32011 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))))))).
32012 exact (fun (x: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))))).
32015 (* constant 6414 *)
32016 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))))))).
32017 exact (fun (x: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)))))).
32020 (* constant 6415 *)
32021 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)))))).
32022 exact (fun (x: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)))))).
32025 (* constant 6416 *)
32026 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)))))))).
32027 exact (fun (x: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)))))).
32030 (* constant 6417 *)
32031 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))))))))))).
32032 exact (fun (x:l_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))))))))))).
32035 (* constant 6418 *)
32036 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)))))))))))).
32037 exact (fun (x:l_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))))))))))).
32040 (* constant 6419 *)
32041 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)))))))))))).
32042 exact (fun (x:l_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))))))))))))).
32045 (* constant 6420 *)
32046 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)))))))))))).
32047 exact (fun (x:l_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)))))))))))).
32050 (* constant 6421 *)
32051 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)))))))))))).
32052 exact (fun (x:l_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))))))))))).
32055 (* constant 6422 *)
32056 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)))))))))))).
32057 exact (fun (x:l_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))))))))))).
32060 (* constant 6423 *)
32061 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)))))))))))).
32062 exact (fun (x:l_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))))))))))))).
32065 (* constant 6424 *)
32066 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)))))))))))).
32067 exact (fun (x:l_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)))))))))))).
32070 (* constant 6425 *)
32071 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))))))))))).
32072 exact (fun (x:l_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))))))))))).
32075 (* constant 6426 *)
32076 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)))))))).
32077 exact (fun (x:l_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))))))).
32080 (* constant 6427 *)
32081 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))))))))).
32082 exact (fun (x: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)))))))).
32085 (* constant 6428 *)
32086 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)))))).
32087 exact (fun (x: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)))))).
32090 (* constant 6429 *)
32091 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).
32092 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)).
32095 (* constant 6430 *)
32096 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.
32097 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)).
32100 (* constant 6431 *)
32101 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)))))).
32102 exact (fun (x: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))))))).
32105 (* constant 6432 *)
32106 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)))))).
32107 exact (fun (x: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))))))).
32110 (* constant 6433 *)
32111 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)))))).
32112 exact (fun (x: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))))))).
32115 (* constant 6434 *)
32116 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)))))).
32117 exact (fun (x: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)))))))).
32120 (* constant 6435 *)
32121 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)))))).
32122 exact (fun (x: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))))))).
32125 (* constant 6436 *)
32126 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)))))).
32127 exact (fun (x: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))))))).
32130 (* constant 6437 *)
32131 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))))).
32132 exact (fun (x: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)))))).
32135 (* constant 6438 *)
32136 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)).
32137 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).
32140 (* constant 6439 *)
32141 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).
32142 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)).
32145 (* constant 6440 *)
32146 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))).
32147 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).
32150 (* constant 6441 *)
32151 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)).
32152 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))).
32155 (* constant 6442 *)
32156 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)).
32157 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)).
32160 (* constant 6443 *)
32161 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))).
32162 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)).
32165 (* constant 6444 *)
32166 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).
32167 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))).
32170 (* constant 6445 *)
32171 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).
32172 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)).
32175 (* constant 6446 *)
32176 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).
32177 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)).
32180 (* constant 6447 *)
32181 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)))))).
32182 exact (fun (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)))))).
32185 (* constant 6448 *)
32186 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)))))).
32187 exact (fun (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)))))).
32190 (* constant 6449 *)
32191 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)))))).
32192 exact (fun (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))))))).
32195 (* constant 6450 *)
32196 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)))))).
32197 exact (fun (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)))))).
32200 (* constant 6451 *)
32201 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)))))).
32202 exact (fun (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)))))).
32205 (* constant 6452 *)
32206 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)))))).
32207 exact (fun (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)))))).
32210 (* constant 6453 *)
32211 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)).
32212 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)).
32215 (* constant 6454 *)
32216 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))).
32217 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))).
32220 (* constant 6455 *)
32221 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)))).
32222 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)))).
32225 (* constant 6456 *)
32226 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)).
32227 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))).
32230 (* constant 6457 *)
32231 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).
32232 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)).
32235 (* constant 6458 *)
32236 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)).
32237 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)))).
32240 (* constant 6459 *)
32241 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))).
32242 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))))).
32245 (* constant 6460 *)
32246 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))).
32247 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))).
32250 (* constant 6461 *)
32251 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)))).
32252 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)))).
32255 (* constant 6462 *)
32256 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)))).
32257 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))).
32260 (* constant 6463 *)
32261 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)))).
32262 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))).
32265 (* constant 6464 *)
32266 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)).
32267 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)).
32270 (* constant 6465 *)
32271 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))).
32272 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))).
32275 (* constant 6466 *)
32276 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))).
32277 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))).
32280 (* constant 6467 *)
32281 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))).
32282 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))).
32285 (* constant 6468 *)
32286 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))).
32287 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))).
32290 (* constant 6469 *)
32291 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))).
32292 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))).
32295 (* constant 6470 *)
32296 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))).
32297 exact (fun (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))))).
32300 (* constant 6471 *)
32301 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))).
32302 exact (fun (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))).
32305 (* constant 6472 *)
32306 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)).
32307 exact (fun (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))).
32310 (* constant 6473 *)
32311 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)))))).
32312 exact (fun (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)))).
32315 (* constant 6474 *)
32316 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)))))).
32317 exact (fun (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)))))).
32320 (* constant 6475 *)
32321 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)))))).
32322 exact (fun (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))))).
32325 (* constant 6476 *)
32326 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)))))).
32327 exact (fun (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))))).
32330 (* constant 6477 *)
32331 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)))))).
32332 exact (fun (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))))).
32335 (* constant 6478 *)
32336 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))))).
32337 exact (fun (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))))).
32340 (* constant 6479 *)
32341 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))))).
32342 exact (fun (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))))).
32345 (* constant 6480 *)
32346 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))).
32347 exact (fun (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))).
32350 (* constant 6481 *)
32351 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))))).
32352 exact (fun (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)))))))).
32355 (* constant 6482 *)
32356 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)))))).
32357 exact (fun (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)))))).
32360 (* constant 6483 *)
32361 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)))))).
32362 exact (fun (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))))))).
32365 (* constant 6484 *)
32366 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)))))))).
32367 exact (fun (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))))))).
32370 (* constant 6485 *)
32371 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)))))).
32372 exact (fun (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))))))).
32375 (* constant 6486 *)
32376 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)))))))).
32377 exact (fun (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)))))))).
32380 (* constant 6487 *)
32381 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))))))))).
32382 exact (fun (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)))))))).
32385 (* constant 6488 *)
32386 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)))))).
32387 exact (fun (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))))))).
32390 (* constant 6489 *)
32391 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)))))).
32392 exact (fun (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))))))).
32395 (* constant 6490 *)
32396 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)))))).
32397 exact (fun (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))))))).
32400 (* constant 6491 *)
32401 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))))))).
32402 exact (fun (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)))))).
32405 (* constant 6492 *)
32406 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)))))))).
32407 exact (fun (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))))))).
32410 (* constant 6493 *)
32411 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)))))))).
32412 exact (fun (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))))))).
32415 (* constant 6494 *)
32416 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)))))))).
32417 exact (fun (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))))))).
32420 (* constant 6495 *)
32421 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)))))))).
32422 exact (fun (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))))))).
32425 (* constant 6496 *)
32426 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)))))))).
32427 exact (fun (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)))))).
32430 (* constant 6497 *)
32431 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)))))))).
32432 exact (fun (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)))))).
32435 (* constant 6498 *)
32436 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)))))))).
32437 exact (fun (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)))))).
32440 (* constant 6499 *)
32441 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)))))))).
32442 exact (fun (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))))))).
32445 (* constant 6500 *)
32446 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))))))))).
32447 exact (fun (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))))))).
32450 (* constant 6501 *)
32451 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))))))))).
32452 exact (fun (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)))))))).
32455 (* constant 6502 *)
32456 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))))))))).
32457 exact (fun (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))))))).
32460 (* constant 6503 *)
32461 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)))))).
32462 exact (fun (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))))))).
32465 (* constant 6504 *)
32466 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)))))).
32467 exact (fun (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))))))).
32470 (* constant 6505 *)
32471 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)))))).
32472 exact (fun (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))))))).
32475 (* constant 6506 *)
32476 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)))))).
32477 exact (fun (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))))))).
32480 (* constant 6507 *)
32481 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)))))))).
32482 exact (fun (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)))))).
32485 (* constant 6508 *)
32486 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)))))))).
32487 exact (fun (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))))))).
32490 (* constant 6509 *)
32491 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)))))))).
32492 exact (fun (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))))))).
32495 (* constant 6510 *)
32496 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))))))))).
32497 exact (fun (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))))))).
32500 (* constant 6511 *)
32501 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)))))))).
32502 exact (fun (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))))))).
32505 (* constant 6512 *)
32506 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)))))))).
32507 exact (fun (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))))))).
32510 (* constant 6513 *)
32511 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)))))).
32512 exact (fun (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))))))).
32515 (* constant 6514 *)
32516 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)))))))).
32517 exact (fun (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)))))).
32520 (* constant 6515 *)
32521 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)).
32522 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).
32525 (* constant 6516 *)
32526 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))).
32527 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))).
32530 (* constant 6517 *)
32531 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)).
32532 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))).
32535 (* constant 6518 *)
32536 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))).
32537 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))).
32540 (* constant 6519 *)
32541 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))))).
32542 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)))).
32545 (* constant 6520 *)
32546 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)))))).
32547 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))).
32550 (* constant 6521 *)
32551 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)))).
32552 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)))).
32555 (* constant 6522 *)
32556 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)))).
32557 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))).
32560 (* constant 6523 *)
32561 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)).
32562 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))))).
32565 (* constant 6524 *)
32566 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)).
32567 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))).
32570 (* constant 6525 *)
32571 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)).
32572 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))).
32575 (* constant 6526 *)
32576 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))))))).
32577 exact (fun (x: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))))))).
32580 (* constant 6527 *)
32581 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))))))).
32582 exact (fun (x: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))))))).
32585 (* constant 6528 *)
32586 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)))))))).
32587 exact (fun (x: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)))))))).
32590 (* constant 6529 *)
32591 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))))))).
32592 exact (fun (x: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))))))).
32595 (* constant 6530 *)
32596 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))))))).
32597 exact (fun (x: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))))))).
32600 (* constant 6531 *)
32601 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)))))))).
32602 exact (fun (x: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))))))).
32605 (* constant 6532 *)
32606 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)))))).
32607 exact (fun (x: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)))))))).
32610 (* constant 6533 *)
32611 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))))))).
32612 exact (fun (x: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)))))))).
32615 (* constant 6534 *)
32616 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))))))).
32617 exact (fun (x: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)))))))).
32620 (* constant 6535 *)
32621 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)))))))))).
32622 exact (fun (x: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))))))))).
32625 (* constant 6536 *)
32626 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))))))).
32627 exact (fun (x: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))))))))).
32630 (* constant 6537 *)
32631 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))))))))).
32632 exact (fun (x: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)))))))).
32635 (* constant 6538 *)
32636 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)))))))))).
32637 exact (fun (x: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))))))))).
32640 (* constant 6539 *)
32641 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)))))))).
32642 exact (fun (x: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))))))))).
32645 (* constant 6540 *)
32646 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)))))))).
32647 exact (fun (x: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)))))))).
32650 (* constant 6541 *)
32651 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))))))))).
32652 exact (fun (x: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)))))))).
32655 (* constant 6542 *)
32656 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))))))))).
32657 exact (fun (x: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)))))))).
32660 (* constant 6543 *)
32661 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))))))))).
32662 exact (fun (x: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))))))))).
32665 (* constant 6544 *)
32666 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))))))))))).
32667 exact (fun (x: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)))))))).
32670 (* constant 6545 *)
32671 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)))))))))).
32672 exact (fun (x: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)))))))).
32675 (* constant 6546 *)
32676 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))))))).
32677 exact (fun (x: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)))))))).
32680 (* constant 6547 *)
32681 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))))))).
32682 exact (fun (x: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))))))).
32685 (* constant 6548 *)
32686 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))))))))).
32687 exact (fun (x: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))))))).
32690 (* constant 6549 *)
32691 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)))))).
32692 exact (fun (x: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)))))).
32695 (* constant 6550 *)
32696 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)))))).
32697 exact (fun (x: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)))))).
32700 (* constant 6551 *)
32701 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)))))).
32702 exact (fun (x: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))))))).
32705 (* constant 6552 *)
32706 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))))))).
32707 exact (fun (x: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)))))).
32710 (* constant 6553 *)
32711 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))))))).
32712 exact (fun (x: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)))))).
32715 (* constant 6554 *)
32716 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))))))).
32717 exact (fun (x: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))))))).
32720 (* constant 6555 *)
32721 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))))))))))).
32722 exact (fun (x: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))))))))))).
32725 (* constant 6556 *)
32726 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)))))))))))).
32727 exact (fun (x: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)))))))))).
32730 (* constant 6557 *)
32731 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))))))))))).
32732 exact (fun (x: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))))))))).
32735 (* constant 6558 *)
32736 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))))))))))).
32737 exact (fun (x: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))))))))).
32740 (* constant 6559 *)
32741 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))))))))))).
32742 exact (fun (x: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)))))))))).
32745 (* constant 6560 *)
32746 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))))))))))).
32747 exact (fun (x: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)))))))))).
32750 (* constant 6561 *)
32751 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)))))))))))).
32752 exact (fun (x: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))))))))))).
32755 (* constant 6562 *)
32756 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)))))))))))).
32757 exact (fun (x: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)))))))))).
32760 (* constant 6563 *)
32761 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))))))))))).
32762 exact (fun (x: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)))))))))).
32765 (* constant 6564 *)
32766 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)))))))))).
32767 exact (fun (x: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)))))))))))).
32770 (* constant 6565 *)
32771 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))))))))))).
32772 exact (fun (x: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)))))))))).
32775 (* constant 6566 *)
32776 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))))))))))).
32777 exact (fun (x: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)))))))))).
32780 (* constant 6567 *)
32781 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))))))))))).
32782 exact (fun (x: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)))))))))).
32785 (* constant 6568 *)
32786 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))))))))).
32787 exact (fun (x: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)))))))))).
32790 (* constant 6569 *)
32791 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))))))))).
32792 exact (fun (x: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)))))))))).
32795 (* constant 6570 *)
32796 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))))))))).
32797 exact (fun (x: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))))))))))).
32800 (* constant 6571 *)
32801 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)))))))))))).
32802 exact (fun (x: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)))))))))).
32805 (* constant 6572 *)
32806 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))))))))).
32807 exact (fun (x: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)))))))))).
32810 (* constant 6573 *)
32811 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))))))))))).
32812 exact (fun (x: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)))))))))).
32815 (* constant 6574 *)
32816 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))))))))))).
32817 exact (fun (x: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)))))))))).
32820 (* constant 6575 *)
32821 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))))))))))).
32822 exact (fun (x: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))))))))))).
32825 (* constant 6576 *)
32826 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))))))))).
32827 exact (fun (x: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)))))))))).
32830 (* constant 6577 *)
32831 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))))))))))).
32832 exact (fun (x: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))))))))).
32835 (* constant 6578 *)
32836 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))))))))).
32837 exact (fun (x: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)))))))))).
32840 (* constant 6579 *)
32841 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)))))))))))).
32842 exact (fun (x: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)))))))))).
32845 (* constant 6580 *)
32846 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))))))))))).
32847 exact (fun (x: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)))))))))).
32850 (* constant 6581 *)
32851 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))))))))))).
32852 exact (fun (x: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)))))))))).
32855 (* constant 6582 *)
32856 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))))))))))).
32857 exact (fun (x: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)))))))))).
32860 (* constant 6583 *)
32861 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)))))))))))).
32862 exact (fun (x: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)))))))))).
32865 (* constant 6584 *)
32866 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))))))))))).
32867 exact (fun (x: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))))))))))).
32870 (* constant 6585 *)
32871 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))))))))))))).
32872 exact (fun (x: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))))))))))).
32875 (* constant 6586 *)
32876 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)))))))))))))).
32877 exact (fun (x: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))))))))))).
32880 (* constant 6587 *)
32881 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)))))))))))).
32882 exact (fun (x: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))))))))))).
32885 (* constant 6588 *)
32886 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)))))))))))).
32887 exact (fun (x: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))))))))))).
32890 (* constant 6589 *)
32891 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))))))))))))).
32892 exact (fun (x: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))))))))))).
32895 (* constant 6590 *)
32896 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)))))))))))))).
32897 exact (fun (x: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))))))))))).
32900 (* constant 6591 *)
32901 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))))))))))))).
32902 exact (fun (x: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))))))))))).
32905 (* constant 6592 *)
32906 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)))))))))))).
32907 exact (fun (x: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))))))))))).
32910 (* constant 6593 *)
32911 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)))))))))).
32912 exact (fun (x: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))))))))))).
32915 (* constant 6594 *)
32916 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)))))))))))).
32917 exact (fun (x: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))))))))))).
32920 (* constant 6595 *)
32921 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)))))))))))).
32922 exact (fun (x: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))))))))))).
32925 (* constant 6596 *)
32926 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)))))))))))).
32927 exact (fun (x: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))))))))))).
32930 (* constant 6597 *)
32931 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)))))))))).
32932 exact (fun (x: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))))))))))).
32935 (* constant 6598 *)
32936 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)))))))))))).
32937 exact (fun (x: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))))))))))).
32940 (* constant 6599 *)
32941 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)))))))))))).
32942 exact (fun (x: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))))))))))).
32945 (* constant 6600 *)
32946 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)))))))))))).
32947 exact (fun (x: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)))))))))))))).
32950 (* constant 6601 *)
32951 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))))))))))).
32952 exact (fun (x: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)))))))))))).
32955 (* constant 6602 *)
32956 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)))))))))))).
32957 exact (fun (x: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))))))))))).
32960 (* constant 6603 *)
32961 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)))))))))).
32962 exact (fun (x: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))))))))))).
32965 (* constant 6604 *)
32966 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)))))))))))).
32967 exact (fun (x: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))))))))))).
32970 (* constant 6605 *)
32971 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)))))))))).
32972 exact (fun (x: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))))))))))).
32975 (* constant 6606 *)
32976 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)))))))))).
32977 exact (fun (x: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))))))))))).
32980 (* constant 6607 *)
32981 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)))))))))))).
32982 exact (fun (x: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)))))))))))).
32985 (* constant 6608 *)
32986 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)))))))))).
32987 exact (fun (x: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))))))))))).
32990 (* constant 6609 *)
32991 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))))))))))).
32992 exact (fun (x: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))))))))))).
32995 (* constant 6610 *)
32996 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)))))))))))).
32997 exact (fun (x: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))))))))))).
33000 (* constant 6611 *)
33001 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))))))))))))).
33002 exact (fun (x: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))))))))))).
33005 (* constant 6612 *)
33006 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))))))))))).
33007 exact (fun (x: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)))))))))).
33010 (* constant 6613 *)
33011 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))))))))))).
33012 exact (fun (x: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))))))))))).
33015 (* constant 6614 *)
33016 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)))))))))))).
33017 exact (fun (x: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))))))))))).
33020 (* constant 6615 *)
33021 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))))))))))))).
33022 exact (fun (x: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))))))))))).
33025 (* constant 6616 *)
33026 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)))))))))))).
33027 exact (fun (x: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))))))))))).
33030 (* constant 6617 *)
33031 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)))))))))))).
33032 exact (fun (x: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))))))))))).
33035 (* constant 6618 *)
33036 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))))))))))).
33037 exact (fun (x: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)))))))))).
33040 (* constant 6619 *)
33041 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)))))))))))).
33042 exact (fun (x: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))))))))))).
33045 (* constant 6620 *)
33046 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)))))))))))).
33047 exact (fun (x: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))))))))))).
33050 (* constant 6621 *)
33051 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)))))))))).
33052 exact (fun (x: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))))))))))).
33055 (* constant 6622 *)
33056 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)))))))))).
33057 exact (fun (x: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))))))))))).
33060 (* constant 6623 *)
33061 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)))))))))).
33062 exact (fun (x: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))))))))))).
33065 (* constant 6624 *)
33066 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)))))))))))).
33067 exact (fun (x: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)))))))))))).
33070 (* constant 6625 *)
33071 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)))))))))))).
33072 exact (fun (x: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))))))))))).
33075 (* constant 6626 *)
33076 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)))))))))))).
33077 exact (fun (x: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))))))))))).
33080 (* constant 6627 *)
33081 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)))))))))))).
33082 exact (fun (x: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))))))))))).
33085 (* constant 6628 *)
33086 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)))))))))))).
33087 exact (fun (x: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))))))))))).
33090 (* constant 6629 *)
33091 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)))))))))))).
33092 exact (fun (x: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))))))))))).
33095 (* constant 6630 *)
33096 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)))))))))))).
33097 exact (fun (x: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)))))))))).
33100 (* constant 6631 *)
33101 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))))))))))).
33102 exact (fun (x: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)))))))))))).
33105 (* constant 6632 *)
33106 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))))))))))).
33107 exact (fun (x: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)))))))))))).
33110 (* constant 6633 *)
33111 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))))))))))).
33112 exact (fun (x: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))))))))))).
33115 (* constant 6634 *)
33116 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)))))))))))).
33117 exact (fun (x: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))))))))))).
33120 (* constant 6635 *)
33121 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)))))))))).
33122 exact (fun (x: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)))))))))))).
33125 (* constant 6636 *)
33126 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)))))))))))).
33127 exact (fun (x: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))))))))))).
33130 (* constant 6637 *)
33131 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)))))))))))).
33132 exact (fun (x: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))))))))))).
33135 (* constant 6638 *)
33136 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)))))))))))).
33137 exact (fun (x: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)))))))))))).
33140 (* constant 6639 *)
33141 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)))))))))).
33142 exact (fun (x: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))))))))))).
33145 (* constant 6640 *)
33146 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))))))))).
33147 exact (fun (x: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)))))))))).
33150 (* constant 6641 *)
33151 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))))))))).
33152 exact (fun (x: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)))))))).
33155 (* constant 6642 *)
33156 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))))))))).
33157 exact (fun (x: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))))))).
33160 (* constant 6643 *)
33161 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)))))))).
33162 exact (fun (x: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)))))))).
33165 (* constant 6644 *)
33166 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)))))))).
33167 exact (fun (x: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)))))))).
33170 (* constant 6645 *)
33171 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))))))))).
33172 exact (fun (x: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)))))))).
33175 (* constant 6646 *)
33176 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)))))))))).
33177 exact (fun (x: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))))))))).
33180 (* constant 6647 *)
33181 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)))))))))).
33182 exact (fun (x: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))))))))).
33185 (* constant 6648 *)
33186 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))))))))))).
33187 exact (fun (x: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))))))))).
33190 (* constant 6649 *)
33191 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))))))))).
33192 exact (fun (x: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)))))))).
33195 (* constant 6650 *)
33196 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))))))))).
33197 exact (fun (x: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)))))))))).
33200 (* constant 6651 *)
33201 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))))))))))).
33202 exact (fun (x: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))))))))).
33205 (* constant 6652 *)
33206 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)))))))))).
33207 exact (fun (x: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)))))))))).
33210 (* constant 6653 *)
33211 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))))))))).
33212 exact (fun (x: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))))))))).
33215 (* constant 6654 *)
33216 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)))))))))).
33217 exact (fun (x: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))))))))).
33220 (* constant 6655 *)
33221 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)))))))).
33222 exact (fun (x: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))))))))).
33225 (* constant 6656 *)
33226 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)))))))))).
33227 exact (fun (x: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)))))))).
33230 (* constant 6657 *)
33231 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)))))))).
33232 exact (fun (x: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))))))))).
33235 (* constant 6658 *)
33236 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))))))))).
33237 exact (fun (x: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))))))))).
33240 (* constant 6659 *)
33241 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)))))))).
33242 exact (fun (x: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))))))))).
33245 (* constant 6660 *)
33246 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)))))))).
33247 exact (fun (x: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))))))))).
33250 (* constant 6661 *)
33251 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))))))).
33252 exact (fun (x: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)))))))).
33255 (* constant 6662 *)
33256 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)))))))).
33257 exact (fun (x: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))))))).
33260 (* constant 6663 *)
33261 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))))))).
33262 exact (fun (x: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)))))).
33265 (* constant 6664 *)
33266 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))))))).
33267 exact (fun (x: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)))))).
33270 (* constant 6665 *)
33271 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)))))))).
33272 exact (fun (x: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)))))).
33275 (* constant 6666 *)
33276 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)))))))).
33277 exact (fun (x: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)))))).
33280 (* constant 6667 *)
33281 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))))))).
33282 exact (fun (x: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)))))).
33285 (* constant 6668 *)
33286 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)))))))).
33287 exact (fun (x: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))))))).
33290 (* constant 6669 *)
33291 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))))))).
33292 exact (fun (x: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))))))).
33295 (* constant 6670 *)
33296 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)))))))).
33297 exact (fun (x: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))))))).
33300 (* constant 6671 *)
33301 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))))))))).
33302 exact (fun (x: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)))))))).
33305 (* constant 6672 *)
33306 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)))))))).
33307 exact (fun (x: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))))))).
33310 (* constant 6673 *)
33311 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)))))).
33312 exact (fun (x: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)))))).
33315 (* constant 6674 *)
33316 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)))))))).
33317 exact (fun (x: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))))))).
33320 (* constant 6675 *)
33321 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)))))))).
33322 exact (fun (x: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))))))).
33325 (* constant 6676 *)
33326 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)))))).
33327 exact (fun (x: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)))))).
33330 (* constant 6677 *)
33331 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)))))))).
33332 exact (fun (x: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))))))).
33335 (* constant 6678 *)
33336 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))))).
33337 exact (fun (x: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)))).
33340 (* constant 6679 *)
33341 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)))).
33342 exact (fun (x: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))))).
33345 (* constant 6680 *)
33346 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))))).
33347 exact (fun (x: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)))).
33350 (* constant 6681 *)
33351 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))))).
33352 exact (fun (x: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)))).
33355 (* constant 6682 *)
33356 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)))))).
33357 exact (fun (x: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)))).
33360 (* constant 6683 *)
33361 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)))).
33362 exact (fun (x: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)))).
33365 (* constant 6684 *)
33366 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)))).
33367 exact (fun (x: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)))).
33370 (* constant 6685 *)
33371 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)))))).
33372 exact (fun (x: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)))))).
33375 (* constant 6686 *)
33376 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)))))).
33377 exact (fun (x: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))))).
33380 (* constant 6687 *)
33381 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)))))).
33382 exact (fun (x: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))))).
33385 (* constant 6688 *)
33386 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)))))).
33387 exact (fun (x: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)))).
33390 (* constant 6689 *)
33391 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))))).
33392 exact (fun (x: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))))).
33395 (* constant 6690 *)
33396 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)))))).
33397 exact (fun (x: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)))).
33400 (* constant 6691 *)
33401 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)))))).
33402 exact (fun (x: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))))).
33405 (* constant 6692 *)
33406 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)))))).
33407 exact (fun (x: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))))).
33410 (* constant 6693 *)
33411 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)))))).
33412 exact (fun (x: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)))).
33415 (* constant 6694 *)
33416 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))))))).
33417 exact (fun (x: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))))))).
33420 (* constant 6695 *)
33421 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))))))).
33422 exact (fun (x: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))))))).
33425 (* constant 6696 *)
33426 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)))))))).
33427 exact (fun (x: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)))))))).
33430 (* constant 6697 *)
33431 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))))))).
33432 exact (fun (x: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))))))).
33435 (* constant 6698 *)
33436 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)))))))).
33437 exact (fun (x: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))))))).
33440 (* constant 6699 *)
33441 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))))).
33442 exact (fun (x: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)))))).
33445 (* constant 6700 *)
33446 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))))).
33447 exact (fun (x: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)))))).
33450 (* constant 6701 *)
33451 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))))))).
33452 exact (fun (x: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)))))).
33455 (* constant 6702 *)
33456 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))))).
33457 exact (fun (x: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))))).
33460 (* constant 6703 *)
33461 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))))).
33462 exact (fun (x: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)))))).
33465 (* constant 6704 *)
33466 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))))).
33467 exact (fun (x: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))))).
33470 (* constant 6705 *)
33471 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)))))).
33472 exact (fun (x: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))))).
33475 (* constant 6706 *)
33476 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)))))).
33477 exact (fun (x: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)))))))).
33480 (* constant 6707 *)
33481 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))))))).
33482 exact (fun (x: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))))))).
33485 (* constant 6708 *)
33486 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))))))).
33487 exact (fun (x: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)))))))).
33490 (* constant 6709 *)
33491 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))))))).
33492 exact (fun (x: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))))))).
33495 (* constant 6710 *)
33496 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))))))).
33497 exact (fun (x: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)))))))).
33500 (* constant 6711 *)
33501 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))))).
33502 exact (fun (x: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)))).
33505 (* constant 6712 *)
33506 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)))).
33507 exact (fun (x: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))))).
33510 (* constant 6713 *)
33511 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))))).
33512 exact (fun (x: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))))).
33515 (* constant 6714 *)
33516 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)))))).
33517 exact (fun (x: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)))))).
33520 (* constant 6715 *)
33521 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))))))).
33522 exact (fun (x: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))))))).
33525 (* constant 6716 *)
33526 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)))))))).
33527 exact (fun (x: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))))).
33530 (* constant 6717 *)
33531 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))))))).
33532 exact (fun (x: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)))))).
33535 (* constant 6718 *)
33536 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))))).
33537 exact (fun (x: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))))))).
33540 (* constant 6719 *)
33541 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))))))).
33542 exact (fun (x: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)))))).
33545 (* constant 6720 *)
33546 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))))).
33547 exact (fun (x: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))))).
33550 (* constant 6721 *)
33551 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)))))).
33552 exact (fun (x: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))))).
33555 (* constant 6722 *)
33556 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)))).
33557 exact (fun (x: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))))).
33560 (* constant 6723 *)
33561 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)))))).
33562 exact (fun (x: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)))))).
33565 (* constant 6724 *)
33566 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)))))).
33567 exact (fun (x: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)))))).
33570 (* constant 6725 *)
33571 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)))))))).
33572 exact (fun (x: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))))))).
33575 (* constant 6726 *)
33576 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)))))))).
33577 exact (fun (x: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))))))).
33580 (* constant 6727 *)
33581 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))))))).
33582 exact (fun (x: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))))))).
33585 (* constant 6728 *)
33586 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)))))))).
33587 exact (fun (x: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))))))).
33590 (* constant 6729 *)
33591 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)))))))).
33592 exact (fun (x: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))))))).
33595 (* constant 6730 *)
33596 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)))))))).
33597 exact (fun (x: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))))))).
33600 (* constant 6731 *)
33601 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))))))))).
33602 exact (fun (x: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))))))).
33605 (* constant 6732 *)
33606 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)))))))).
33607 exact (fun (x: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)))))).
33610 (* constant 6733 *)
33611 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))))))))).
33612 exact (fun (x: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))))))).
33615 (* constant 6734 *)
33616 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))))))).
33617 exact (fun (x: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))))))).
33620 (* constant 6735 *)
33621 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)))))))).
33622 exact (fun (x: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))))))).
33625 (* constant 6736 *)
33626 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)))))))).
33627 exact (fun (x: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)))))))).
33630 (* constant 6737 *)
33631 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)))))))).
33632 exact (fun (x: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)))))))).
33635 (* constant 6738 *)
33636 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))))))))).
33637 exact (fun (x: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))))))).
33640 (* constant 6739 *)
33641 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)))))))).
33642 exact (fun (x: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))))))).
33645 (* constant 6740 *)
33646 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))))))))).
33647 exact (fun (x: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)))))))).
33650 (* constant 6741 *)
33651 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))))))))).
33652 exact (fun (x: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)))))))).
33655 (* constant 6742 *)
33656 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))))))).
33657 exact (fun (x: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))))))).
33660 (* constant 6743 *)
33661 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))))))).
33662 exact (fun (x: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))))))).
33665 (* constant 6744 *)
33666 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))))).
33667 exact (fun (x: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))))).
33670 (* constant 6745 *)
33671 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)))))))).
33672 exact (fun (x: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)))))))).
33675 (* constant 6746 *)
33676 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)))))))).
33677 exact (fun (x: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))))))))).
33680 (* constant 6747 *)
33681 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)))))))))).
33682 exact (fun (x: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))))))))).
33685 (* constant 6748 *)
33686 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)))))))).
33687 exact (fun (x: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))))))))).
33690 (* constant 6749 *)
33691 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)))))).
33692 exact (fun (x: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))))))).
33695 (* constant 6750 *)
33696 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))))))))).
33697 exact (fun (x: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))))))))).
33700 (* constant 6751 *)
33701 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))))))))))).
33702 exact (fun (x: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)))))))))).
33705 (* constant 6752 *)
33706 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))))))))))).
33707 exact (fun (x: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))))))))).
33710 (* constant 6753 *)
33711 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))))))))).
33712 exact (fun (x: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))))))))).
33715 (* constant 6754 *)
33716 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)))))))))).
33717 exact (fun (x: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))))))))).
33720 (* constant 6755 *)
33721 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)))))))).
33722 exact (fun (x: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))))))))).
33725 (* constant 6756 *)
33726 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)))))))).
33727 exact (fun (x: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)))))))).
33730 (* constant 6757 *)
33731 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)))))))).
33732 exact (fun (x: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)))))))).
33735 (* constant 6758 *)
33736 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)))))))).
33737 exact (fun (x: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))))))))).
33740 (* constant 6759 *)
33741 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)))))))).
33742 exact (fun (x: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))))))))).
33745 (* constant 6760 *)
33746 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)))))))).
33747 exact (fun (x: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)))))))).
33750 (* constant 6761 *)
33751 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))))))))).
33752 exact (fun (x: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)))))))).
33755 (* constant 6762 *)
33756 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))))))))))).
33757 exact (fun (x: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)))))))).
33760 (* constant 6763 *)
33761 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)))))))).
33762 exact (fun (x: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))))))))).
33765 (* constant 6764 *)
33766 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))))))))).
33767 exact (fun (x: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))))))))).
33770 (* constant 6765 *)
33771 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))))))))).
33772 exact (fun (x: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))))))))).
33775 (* constant 6766 *)
33776 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)))))))))).
33777 exact (fun (x: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))))))))).
33780 (* constant 6767 *)
33781 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))))))))))).
33782 exact (fun (x: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))))))))).
33785 (* constant 6768 *)
33786 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)))))))))).
33787 exact (fun (x: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))))))))).
33790 (* constant 6769 *)
33791 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)))))))))).
33792 exact (fun (x: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))))))))).
33795 (* constant 6770 *)
33796 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)))))))))).
33797 exact (fun (x: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))))))))).
33800 (* constant 6771 *)
33801 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)))))))).
33802 exact (fun (x: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)))))))).
33805 (* constant 6772 *)
33806 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)))))))))).
33807 exact (fun (x: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))))))))).
33810 (* constant 6773 *)
33811 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)))))))).
33812 exact (fun (x: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))))))))).
33815 (* constant 6774 *)
33816 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))))))))))).
33817 exact (fun (x: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))))))))).
33820 (* constant 6775 *)
33821 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)))))))))).
33822 exact (fun (x: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))))))))).
33825 (* constant 6776 *)
33826 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))))))))).
33827 exact (fun (x: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))))))))).
33830 (* constant 6777 *)
33831 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)))))))))).
33832 exact (fun (x: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))))))))).
33835 (* constant 6778 *)
33836 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)))))))))).
33837 exact (fun (x: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))))))))).
33840 (* constant 6779 *)
33841 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))))))))).
33842 exact (fun (x: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))))))))))).
33845 (* constant 6780 *)
33846 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))))))))))).
33847 exact (fun (x: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)))))))))).
33850 (* constant 6781 *)
33851 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))))))))))).
33852 exact (fun (x: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)))))))).
33855 (* constant 6782 *)
33856 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)))))))).
33857 exact (fun (x: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)))))))).
33860 (* constant 6783 *)
33861 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)))))))))))).
33862 exact (fun (x: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)))))))).
33865 (* constant 6784 *)
33866 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)))))))))).
33867 exact (fun (x: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)))))))).
33870 (* constant 6785 *)
33871 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)))))))))).
33872 exact (fun (x: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))))))))).
33875 (* constant 6786 *)
33876 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)))))))))).
33877 exact (fun (x: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))))))))).
33880 (* constant 6787 *)
33881 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)))))))))).
33882 exact (fun (x: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))))))))).
33885 (* constant 6788 *)
33886 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)))))))))).
33887 exact (fun (x: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))))))))).
33890 (* constant 6789 *)
33891 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)))))))))).
33892 exact (fun (x: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))))))))).
33895 (* constant 6790 *)
33896 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)))))))).
33897 exact (fun (x: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))))))))).
33900 (* constant 6791 *)
33901 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))))))).
33902 exact (fun (x: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)))))))).
33905 (* constant 6792 *)
33906 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)))))))).
33907 exact (fun (x: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))))))).
33910 (* constant 6793 *)
33911 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))).
33912 exact (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))))).
33915 (* constant 6794 *)
33916 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)))).
33917 exact (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))).
33920 (* constant 6795 *)
33921 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))).
33922 exact (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)).
33925 (* constant 6796 *)
33926 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))).
33927 exact (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)))))).
33930 (* constant 6797 *)
33931 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)))).
33932 exact (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))).
33935 (* constant 6798 *)
33936 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))).
33937 exact (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)).
33940 (* constant 6799 *)
33941 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))).
33942 exact (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))))).
33945 (* constant 6800 *)
33946 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)).
33947 exact (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)))).
33950 (* constant 6801 *)
33951 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))).
33952 exact (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)))).
33955 (* constant 6802 *)
33956 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)))).
33957 exact (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))).
33960 (* constant 6803 *)
33961 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))).
33962 exact (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)).
33965 (* constant 6804 *)
33966 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)))).
33967 exact (fun (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))))).
33970 (* constant 6805 *)
33971 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))).
33972 exact (fun (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)))).
33975 (* constant 6806 *)
33976 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)))).
33977 exact (fun (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))))).
33980 (* constant 6807 *)
33981 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))))).
33982 exact (fun (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)))).
33985 (* constant 6808 *)
33986 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)))).
33987 exact (fun (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)))).
33990 (* constant 6809 *)
33991 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)).
33992 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)))).
33995 (* constant 6810 *)
33996 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))).
33997 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)).
34000 (* constant 6811 *)
34001 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)).
34002 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).
34005 (* constant 6812 *)
34006 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)).
34007 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))).
34010 (* constant 6813 *)
34011 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).
34012 exact (fun (r:l_e_st_eq_landau_n_rt_rp_r_real) => l_e_st_eq_landau_n_rt_rp_r_abs r).
34015 (* constant 6814 *)
34016 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)))).
34017 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)).
34020 (* constant 6815 *)
34021 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)))).
34022 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)))).
34025 (* constant 6816 *)
34026 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)))).
34027 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)).
34030 (* constant 6817 *)
34031 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))).
34032 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)).
34035 (* constant 6818 *)
34036 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))).
34037 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)).
34040 (* constant 6819 *)
34041 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))).
34042 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))).
34045 (* constant 6820 *)
34046 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)).
34047 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)).
34050 (* constant 6821 *)
34051 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))).
34052 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)).
34055 (* constant 6822 *)
34056 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).
34057 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).
34060 (* constant 6823 *)
34061 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))).
34062 exact (fun (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)))).
34065 (* constant 6824 *)
34066 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)))).
34067 exact (fun (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))).
34070 (* constant 6825 *)
34071 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).
34072 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))).
34075 (* constant 6826 *)
34076 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).
34077 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).
34080 (* constant 6827 *)
34081 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)).
34082 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).
34085 (* constant 6828 *)
34086 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)).
34087 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)).
34090 (* constant 6829 *)
34091 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))))).
34092 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))))).
34095 (* constant 6830 *)
34096 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)))))).
34097 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))))).
34100 (* constant 6831 *)
34101 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))).
34102 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).
34105 (* constant 6832 *)
34106 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).
34107 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).
34110 (* constant 6833 *)
34111 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)))).
34112 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)).
34115 (* constant 6834 *)
34116 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)).
34117 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)).
34120 (* constant 6835 *)
34121 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).
34122 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)).
34125 (* constant 6836 *)
34126 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)))).
34127 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)))).
34130 (* constant 6837 *)
34131 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))).
34132 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)))).
34135 (* constant 6838 *)
34136 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).
34137 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))).
34140 (* constant 6839 *)
34141 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).
34142 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).
34145 (* constant 6840 *)
34146 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)).
34147 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).
34150 (* constant 6841 *)
34151 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)).
34152 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)).
34155 (* constant 6842 *)
34156 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)))))).
34157 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))))).
34160 (* constant 6843 *)
34161 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))))).
34162 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))))).
34165 (* constant 6844 *)
34166 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))).
34167 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).
34170 (* constant 6845 *)
34171 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).
34172 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).
34175 (* constant 6846 *)
34176 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)))).
34177 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)).
34180 (* constant 6847 *)
34181 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)).
34182 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)).
34185 (* constant 6848 *)
34186 Definition l_e_st_eq_landau_n_rt_rp_r_c_natt : Type.
34187 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)).
34190 (* constant 6849 *)
34191 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).
34192 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).
34195 (* constant 6850 *)
34196 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)).
34197 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).
34200 (* constant 6851 *)
34201 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)))).
34202 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))).
34205 (* constant 6852 *)
34206 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))).
34207 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))).
34210 (* constant 6853 *)
34211 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)).
34212 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)).
34215 (* constant 6854 *)
34216 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)))))).
34217 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))))).
34220 (* constant 6855 *)
34221 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))))).
34222 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))))).
34225 (* constant 6856 *)
34226 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))).
34227 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).
34230 (* constant 6857 *)
34231 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).
34232 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)).
34235 (* constant 6858 *)
34236 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)))).
34237 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)).
34240 (* constant 6859 *)
34241 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)).
34242 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))).
34245 (* constant 6860 *)
34246 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).
34247 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)).
34250 (* constant 6861 *)
34251 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)))).
34252 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)))).
34255 (* constant 6862 *)
34256 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))).
34257 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)))).
34260 (* constant 6863 *)
34261 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).
34262 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)).
34265 (* constant 6864 *)
34266 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)))).
34267 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)))).
34270 (* constant 6865 *)
34271 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))).
34272 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)))).
34275 (* constant 6866 *)
34276 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))).
34277 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)).
34280 (* constant 6867 *)
34281 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))).
34282 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))).
34285 (* constant 6868 *)
34286 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).
34287 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)).
34290 (* constant 6869 *)
34291 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))).
34292 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)).
34295 (* constant 6870 *)
34296 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))).
34297 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))).
34300 (* constant 6871 *)
34301 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).
34302 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)).
34305 (* constant 6872 *)
34306 Definition l_e_st_eq_landau_n_rt_rp_r_c_1t : l_e_st_eq_landau_n_rt_rp_r_c_natt.
34307 exact (l_e_st_eq_landau_n_rt_rp_r_c_ntofn l_e_st_eq_landau_n_1).
34310 (* constant 6873 *)
34311 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).
34312 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))).
34315 (* constant 6874 *)
34316 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)).
34317 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)).
34320 (* constant 6875 *)
34321 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)).
34322 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)).
34325 (* constant 6876 *)
34326 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)).
34327 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).
34330 (* constant 6877 *)
34331 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))))).
34332 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))).
34335 (* constant 6878 *)
34336 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)))).
34337 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)))).
34340 (* constant 6879 *)
34341 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))).
34342 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)))).
34345 (* constant 6880 *)
34346 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))).
34347 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))).
34350 (* constant 6881 *)
34351 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).
34352 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).
34355 (* constant 6882 *)
34356 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).
34357 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))).
34360 (* constant 6883 *)
34361 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)))).
34362 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)))).
34365 (* constant 6884 *)
34366 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))).
34367 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))).
34370 (* constant 6885 *)
34371 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))))).
34372 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))))).
34375 (* constant 6886 *)
34376 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)))))).
34377 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)))))).
34380 (* constant 6887 *)
34381 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))))).
34382 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))))).
34385 (* constant 6888 *)
34386 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)))).
34387 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))))).
34390 (* constant 6889 *)
34391 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)))).
34392 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))).
34395 (* constant 6890 *)
34396 Definition l_e_st_eq_landau_n_rt_rp_r_c_ic : l_e_st_eq_landau_n_rt_rp_r_c_complex.
34397 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).
34400 (* constant 6891 *)
34401 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))).
34402 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).
34405 (* constant 6892 *)
34406 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).
34407 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))).
34410 (* constant 6893 *)
34411 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.
34412 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))).
34415 (* constant 6894 *)
34416 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)).
34417 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).
34420 (* constant 6895 *)
34421 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).
34422 exact (l_e_st_eq_landau_n_rt_rp_r_c_satz298j l_e_st_eq_landau_n_rt_rp_r_1rl).
34425 (* constant 6896 *)
34426 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).
34427 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).
34430 (* constant 6897 *)
34431 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))))).
34432 exact (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)).
34435 (* constant 6898 *)
34436 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)).
34437 exact (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))))).
34440 (* constant 6899 *)
34441 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)).
34442 exact (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))).
34445 (* constant 6900 *)
34446 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))).
34447 exact (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))).
34450 (* constant 6901 *)
34451 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))).
34452 exact (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))).
34455 (* constant 6902 *)
34456 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)))).
34457 exact (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))).
34460 (* constant 6903 *)
34461 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)))).
34462 exact (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)).
34465 (* constant 6904 *)
34466 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))).
34467 exact (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)))).
34470 (* constant 6905 *)
34471 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))).
34472 exact (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))).
34475 (* constant 6906 *)
34476 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)))).
34477 exact (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))).
34480 (* constant 6907 *)
34481 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))).
34482 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))).
34485 (* constant 6908 *)
34486 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).
34487 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)).
34490 (* constant 6909 *)
34491 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)))))).
34492 exact (fun (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)))))).
34495 (* constant 6910 *)
34496 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))))).
34497 exact (fun (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)))))).
34500 (* constant 6911 *)
34501 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))))).
34502 exact (fun (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)))))).