(**************************************************************************) (* ___ *) (* ||M|| *) (* ||A|| A project by Andrea Asperti *) (* ||T|| *) (* ||I|| Developers: *) (* ||T|| The HELM team. *) (* ||A|| http://helm.cs.unibo.it *) (* \ / *) (* \ / This file is distributed under the terms of the *) (* v GNU General Public License Version 2 *) (* *) (**************************************************************************) include "basics/types.ma". include "basics/lists/list.ma". include "basics/logic.ma". include "ASM/Util.ma". lemma eq_rect_Type0_r: ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[0]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p. #A #a #P #p #x0 #p0 @(eq_rect_r ??? p0) assumption. qed. lemma eq_rect_r2: ∀A.∀a,x.∀p:eq ? x a.∀P: ∀x:A. eq ? x a → Type[2]. P a (refl A a) → P x p. #A #a #x #p cases p; #P #H assumption. qed. lemma eq_rect_Type2_r: ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[2]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p. #A #a #P #p #x0 #p0 @(eq_rect_r2 ??? p0) assumption. qed. lemma eq_rect_CProp0_r: ∀A.∀a.∀P: ∀x:A. eq ? x a → CProp[0]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p. #A #a #P #p #x0 #p0 @(eq_rect_r2 ??? p0) assumption. qed. lemma sym_neq : ∀A.∀x,y:A. x ≠ y → y ≠ x. #A #x #y *;#H @nmk #H' /2/; qed. (* Paolo: already in library, generates ambiguity interpretation "logical iff" 'iff x y = (iff x y).*) (* bool *) lemma xorb_neg : ∀a,b. notb (xorb a b) = xorb a (notb b). * * @refl qed. lemma xorb_false : ∀a. xorb a false = a. * @refl qed. lemma xorb_true : ∀a. xorb a true = (¬a). * @refl qed. lemma xorb_comm : ∀a,b. xorb a b = xorb b a. * * @refl qed. lemma xorb_assoc : ∀a,b,c. xorb a (xorb b c) = xorb (xorb a b) c. * * * @refl qed. lemma xorb_lneg : ∀a,b. xorb (¬a) b = (¬xorb a b). * * @refl qed. lemma xorb_rneg : ∀a,b. xorb a (¬b) = (¬xorb a b). * * @refl qed. lemma xorb_inj : ∀a,b,c. xorb a b = xorb a c ↔ b = c. * * * @conj try // normalize try // qed. (* should be proved in nat.ma, but it's not! *) (* Paolo: there is eqb_elim which does something very similar *) lemma eqb_to_Prop : ∀n,m:nat.match eqb n m with [ true ⇒ n = m | false ⇒ n ≠ m ]. #n elim n; [ #m cases m; //; | #n' #IH #m cases m; [ /2/; | #m' whd in match (eqb (S n') (S m')) in ⊢ %; lapply (IH m'); cases (eqb n' m'); /2/; ] ] qed. (* datatypes/list.ma *) theorem nil_append_nil_both: ∀A:Type[0]. ∀l1,l2:list A. l1 @ l2 = [] → l1 = [] ∧ l2 = []. #A #l1 #l2 cases l1 [ cases l2 [ /2/ | normalize #h #t #H destruct ] | cases l2 [ normalize #h #t #H destruct | normalize #h1 #t1 #h2 #h3 #H destruct ] ] qed. (* some useful stuff for quantifiers *) lemma dec_bounded_forall: ∀P:ℕ → Prop.(∀n.(P n) + (¬P n)) → ∀k.(∀n.n < k → P n) + ¬(∀n.n < k → P n). #P #HP_dec #k elim k -k [ %1 #n #Hn @⊥ @(absurd (n < 0) Hn) @not_le_Sn_O | #k #Hind cases Hind [ #Hk cases (HP_dec k) [ #HPk %1 #n #Hn cases (le_to_or_lt_eq … Hn) [ #H @(Hk … (le_S_S_to_le … H)) | #H >(injective_S … H) @HPk ] | #HPk %2 @nmk #Habs @(absurd (P k)) [ @(Habs … (le_n (S k))) | @HPk ] ] | #Hk %2 @nmk #Habs @(absurd (∀n.n(injective_S … H2) @HP ] ] | #HP @(ex_intro … k) @conj [ @le_n | @HP ] ] ] qed. lemma associative_orb : associative ? orb. *** // qed. lemma commutative_orb : commutative ? orb. ** // qed. lemma associative_andb : associative ? andb. *** // qed. lemma commutative_andb : commutative ? andb. ** // qed. lemma notb_false : ∀b.(¬b) = false → b = true. * [#_ % | normalize #EQ destruct] qed. lemma notb_true : ∀b.(¬b) = true → b = false. * [normalize #EQ destruct | #_ %] qed. notation > "Σ 〈 ident x : tyx, ident y : tyy 〉 . P" with precedence 20 for @{'sigma (λ${fresh p}. match ${fresh p} with [mk_Prod (${ident x} : $tyx) (${ident y} : $tyy) ⇒ $P])}. notation > "Σ 〈 ident x, ident y 〉 . P" with precedence 20 for @{'sigma (λ${fresh p}. match ${fresh p} with [mk_Prod ${ident x} ${ident y} ⇒ $P])}. notation > "Σ 〈 ident x : tyx, ident y : tyy, ident z : tyz 〉 . P" with precedence 20 for @{'sigma (λ${fresh p1}. match ${fresh p1} with [mk_Prod ${fresh p2} (${ident z} : $tyz) ⇒ match ${fresh p2} with [mk_Prod (${ident x} : $tyx) (${ident y} : $tyy) ⇒ $P]])}. notation > "Σ 〈 ident x , ident y , ident z 〉 . P" with precedence 20 for @{'sigma (λ${fresh p1}. match ${fresh p1} with [mk_Prod ${fresh p2} ${ident z} ⇒ match ${fresh p2} with [mk_Prod ${ident x} ${ident y} ⇒ $P]])}. notation > "vbox('let' ❬ ident v , ident p ❭ ≝ e 'in' break e')" with precedence 10 for @{ match ${e} with [ mk_DPair ${ident v} ${ident p} ⇒ ${e'} ] }. notation > "hvbox('let' ❬ident x,ident y❭ 'as' ident E ≝ t 'in' s)" with precedence 10 for @{ match $t return λx.x = $t → ? with [ mk_DPair ${ident x} ${ident y} ⇒ λ${ident E}.$s ] (refl ? $t) }.