]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/bacato.ma
one main property of lift closed
[helm.git] / matita / matita / lib / bacato.ma
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic
3     ||A||  Library of Mathematics, developed at the Computer Science
4     ||T||  Department of the University of Bologna, Italy.
5     ||I||
6     ||T||
7     ||A||  This file is distributed under the terms of the
8     \   /  GNU General Public License Version 2
9      \ /
10       V_______________________________________________________________ *)
11
12 include "basics/pts.ma".
13
14 inductive ex3_2 (A0,A1:Type[0]) (P0,P1,P2:A0→A1→Prop) : Prop ≝
15    | ex3_2_intro: ∀x0,x1. (P0 x0 x1)→(P1 x0 x1)→(P2 x0 x1)→(ex3_2 ? ? ? ? ?).
16
17 lemma bacato: ∀A0,A1:Type[0]. ∀P0,P1,P2:A0→A1→Prop. ∀x0,x1.
18               P0 x0 x1 → P1 x0 x1 → P2 x0 x1 →
19               ex3_2 … (λx0,x1. P0 x0 x1) (λx0,x1. P1 x0 x1) (λx0,x1. P2 x0 x1).
20 #A0 #A1 #P0 #P1 #P2 #x0 #x1 #H0 #H1 #H2
21 @ex3_2_intro //