include "basics/lists/list.ma". include "basics/types.ma". include "arithmetics/nat.ma". include "basics/russell.ma". (* let's implement a daemon not used by automation *) inductive DAEMONXXX: Type[0] ≝ K1DAEMONXXX: DAEMONXXX | K2DAEMONXXX: DAEMONXXX. axiom IMPOSSIBLE: K1DAEMONXXX = K2DAEMONXXX. example daemon: False. generalize in match IMPOSSIBLE; #IMPOSSIBLE destruct(IMPOSSIBLE) qed. example not_implemented: False. cases daemon qed. notation "⊥" with precedence 90 for @{ match ? in False with [ ] }. notation "Ⓧ" with precedence 90 for @{ λabs.match abs in False with [ ] }. definition ltb ≝ λm, n: nat. leb (S m) n. definition geb ≝ λm, n: nat. leb n m. definition gtb ≝ λm, n: nat. ltb n m. (* dpm: unless I'm being stupid, this isn't defined in the stdlib? *) let rec eq_nat (n: nat) (m: nat) on n: bool ≝ match n with [ O ⇒ match m with [ O ⇒ true | _ ⇒ false ] | S n' ⇒ match m with [ S m' ⇒ eq_nat n' m' | _ ⇒ false ] ]. let rec forall (A: Type[0]) (f: A → bool) (l: list A) on l ≝ match l with [ nil ⇒ true | cons hd tl ⇒ f hd ∧ forall A f tl ]. let rec prefix (A: Type[0]) (k: nat) (l: list A) on l ≝ match l with [ nil ⇒ [ ] | cons hd tl ⇒ match k with [ O ⇒ [ ] | S k' ⇒ hd :: prefix A k' tl ] ]. let rec fold_left2 (A: Type[0]) (B: Type[0]) (C: Type[0]) (f: A → B → C → A) (accu: A) (left: list B) (right: list C) (proof: |left| = |right|) on left: A ≝ match left return λx. |x| = |right| → A with [ nil ⇒ λnil_prf. match right return λx. |[ ]| = |x| → A with [ nil ⇒ λnil_nil_prf. accu | cons hd tl ⇒ λcons_nil_absrd. ? ] nil_prf | cons hd tl ⇒ λcons_prf. match right return λx. |hd::tl| = |x| → A with [ nil ⇒ λcons_nil_absrd. ? | cons hd' tl' ⇒ λcons_cons_prf. fold_left2 … f (f accu hd hd') tl tl' ? ] cons_prf ] proof. [ 1: normalize in cons_nil_absrd; destruct(cons_nil_absrd) | 2: normalize in cons_nil_absrd; destruct(cons_nil_absrd) | 3: normalize in cons_cons_prf; @injective_S assumption ] qed. let rec remove_n_first_internal (i: nat) (A: Type[0]) (l: list A) (n: nat) on l ≝ match l with [ nil ⇒ [ ] | cons hd tl ⇒ match eq_nat i n with [ true ⇒ l | _ ⇒ remove_n_first_internal (S i) A tl n ] ]. definition remove_n_first ≝ λA: Type[0]. λn: nat. λl: list A. remove_n_first_internal 0 A l n. let rec foldi_from_until_internal (A: Type[0]) (i: nat) (res: ?) (rem: list A) (m: nat) (f: nat → list A → A → list A) on rem ≝ match rem with [ nil ⇒ res | cons e tl ⇒ match geb i m with [ true ⇒ res | _ ⇒ foldi_from_until_internal A (S i) (f i res e) tl m f ] ]. definition foldi_from_until ≝ λA: Type[0]. λn: nat. λm: nat. λf: ?. λa: ?. λl: ?. foldi_from_until_internal A 0 a (remove_n_first A n l) m f. definition foldi_from ≝ λA: Type[0]. λn. λf. λa. λl. foldi_from_until A n (|l|) f a l. definition foldi_until ≝ λA: Type[0]. λm. λf. λa. λl. foldi_from_until A 0 m f a l. definition foldi ≝ λA: Type[0]. λf. λa. λl. foldi_from_until A 0 (|l|) f a l. definition hd_safe ≝ λA: Type[0]. λl: list A. λproof: 0 < |l|. match l return λx. 0 < |x| → A with [ nil ⇒ λnil_absrd. ? | cons hd tl ⇒ λcons_prf. hd ] proof. normalize in nil_absrd; cases(not_le_Sn_O 0) #HYP cases(HYP nil_absrd) qed. definition tail_safe ≝ λA: Type[0]. λl: list A. λproof: 0 < |l|. match l return λx. 0 < |x| → list A with [ nil ⇒ λnil_absrd. ? | cons hd tl ⇒ λcons_prf. tl ] proof. normalize in nil_absrd; cases(not_le_Sn_O 0) #HYP cases(HYP nil_absrd) qed. let rec safe_split (A: Type[0]) (l: list A) (index: nat) (proof: index ≤ |l|) on index ≝ match index return λx. x ≤ |l| → (list A) × (list A) with [ O ⇒ λzero_prf. 〈[], l〉 | S index' ⇒ λsucc_prf. match l return λx. S index' ≤ |x| → (list A) × (list A) with [ nil ⇒ λnil_absrd. ? | cons hd tl ⇒ λcons_prf. let 〈l1, l2〉 ≝ safe_split A tl index' ? in 〈hd :: l1, l2〉 ] succ_prf ] proof. [1: normalize in nil_absrd; cases(not_le_Sn_O index') #HYP cases(HYP nil_absrd) |2: normalize in cons_prf; @le_S_S_to_le assumption ] qed. let rec nth_safe (elt_type: Type[0]) (index: nat) (the_list: list elt_type) (proof: index < | the_list |) on index ≝ match index return λs. s < | the_list | → elt_type with [ O ⇒ match the_list return λt. 0 < | t | → elt_type with [ nil ⇒ λnil_absurd. ? | cons hd tl ⇒ λcons_proof. hd ] | S index' ⇒ match the_list return λt. S index' < | t | → elt_type with [ nil ⇒ λnil_absurd. ? | cons hd tl ⇒ λcons_proof. nth_safe elt_type index' tl ? ] ] proof. [ normalize in nil_absurd; cases (not_le_Sn_O 0) #ABSURD elim (ABSURD nil_absurd) | normalize in nil_absurd; cases (not_le_Sn_O (S index')) #ABSURD elim (ABSURD nil_absurd) | normalize in cons_proof; @le_S_S_to_le assumption ] qed. lemma shift_nth_safe: ∀T,i,l2,l1,K1,K2. nth_safe T i l1 K1 = nth_safe T (i+|l2|) (l2@l1) K2. #T #i #l2 elim l2 normalize [ #l1 #K1 commutative_plus @shift_nth_safe qed. lemma nth_safe_cons: ∀A,hd,tl,l2,j,j_ok,Sj_ok. nth_safe A j (tl@l2) j_ok =nth_safe A (1+j) (hd::tl@l2) Sj_ok. // qed. lemma eq_nth_safe_proof_irrelevant: ∀A,l,i,i_ok,i_ok'. nth_safe A l i i_ok = nth_safe A l i i_ok'. #A #l #i #i_ok #i_ok' % qed. lemma nth_safe_append: ∀A,l,n,n_ok. ∃pre,suff. (pre @ [nth_safe A n l n_ok]) @ suff = l. #A #l elim l normalize [ #n #abs cases (absurd ? abs (not_le_Sn_O ?)) | #hd #tl #IH #n cases n normalize [ #_ %{[]} /2/ | #m #m_ok cases (IH m ?) [ #pre * #suff #EQ %{(hd::pre)} %{suff} p3 in p2; >p4 cases (reduce_strong … tl tl1) normalize #X #H #EQ destruct // ] qed. let rec map2_opt (A: Type[0]) (B: Type[0]) (C: Type[0]) (f: A → B → C) (left: list A) (right: list B) on left ≝ match left with [ nil ⇒ match right with [ nil ⇒ Some ? (nil C) | _ ⇒ None ? ] | cons hd tl ⇒ match right with [ nil ⇒ None ? | cons hd' tl' ⇒ match map2_opt A B C f tl tl' with [ None ⇒ None ? | Some tail ⇒ Some ? (f hd hd' :: tail) ] ] ]. let rec map2 (A: Type[0]) (B: Type[0]) (C: Type[0]) (f: A → B → C) (left: list A) (right: list B) (proof: | left | = | right |) on left ≝ match left return λx. | x | = | right | → list C with [ nil ⇒ match right return λy. | [] | = | y | → list C with [ nil ⇒ λnil_prf. nil C | _ ⇒ λcons_absrd. ? ] | cons hd tl ⇒ match right return λy. | hd::tl | = | y | → list C with [ nil ⇒ λnil_absrd. ? | cons hd' tl' ⇒ λcons_prf. (f hd hd') :: map2 A B C f tl tl' ? ] ] proof. [1: normalize in cons_absrd; destruct(cons_absrd) |2: normalize in nil_absrd; destruct(nil_absrd) |3: normalize in cons_prf; destruct(cons_prf) assumption ] qed. let rec map3 (A: Type[0]) (B: Type[0]) (C: Type[0]) (D: Type[0]) (f: A → B → C → D) (left: list A) (centre: list B) (right: list C) (prflc: |left| = |centre|) (prfcr: |centre| = |right|) on left ≝ match left return λx. |x| = |centre| → list D with [ nil ⇒ λnil_prf. match centre return λx. |x| = |right| → list D with [ nil ⇒ λnil_nil_prf. match right return λx. |nil ?| = |x| → list D with [ nil ⇒ λnil_nil_nil_prf. nil D | cons hd tl ⇒ λcons_nil_nil_absrd. ? ] nil_nil_prf | cons hd tl ⇒ λnil_cons_absrd. ? ] prfcr | cons hd tl ⇒ λcons_prf. match centre return λx. |x| = |right| → list D with [ nil ⇒ λcons_nil_absrd. ? | cons hd' tl' ⇒ λcons_cons_prf. match right return λx. |right| = |x| → |cons ? hd' tl'| = |x| → list D with [ nil ⇒ λrefl_prf. λcons_cons_nil_absrd. ? | cons hd'' tl'' ⇒ λrefl_prf. λcons_cons_cons_prf. (f hd hd' hd'') :: (map3 A B C D f tl tl' tl'' ? ?) ] (refl ? (|right|)) cons_cons_prf ] prfcr ] prflc. [ 1: normalize in cons_nil_nil_absrd; destruct(cons_nil_nil_absrd) | 2: generalize in match nil_cons_absrd; append_nil % | #hd #tl #IH #pre whd in ⊢ (???%); <(foldl_step … H ??) applyS (IH (pre@[hd])) ] qed. (* redirecting to library reverse *) definition rev ≝ reverse. lemma append_length: ∀A: Type[0]. ∀l, r: list A. |(l @ r)| = |l| + |r|. #A #L #R elim L [ % | #HD #TL #IH normalize >IH % ] qed. lemma append_nil: ∀A: Type[0]. ∀l: list A. l @ [ ] = l. #A #L elim L // qed. lemma rev_append: ∀A: Type[0]. ∀l, r: list A. rev A (l @ r) = rev A r @ rev A l. #A #L #R elim L [ normalize >append_nil % | #HD #TL normalize #IH >rev_append_def >rev_append_def >rev_append_def >append_nil normalize >IH @associative_append ] qed. lemma rev_length: ∀A: Type[0]. ∀l: list A. |rev A l| = |l|. #A #L elim L [ % | #HD #TL normalize #IH >rev_append_def >(append_length A (rev A TL) [HD]) normalize /2 by / ] qed. lemma nth_append_first: ∀A:Type[0]. ∀n:nat.∀l1,l2:list A.∀d:A. n < |l1| → nth n A (l1@l2) d = nth n A l1 d. #A #n #l1 #l2 #d generalize in match n; -n; elim l1 [ normalize #k #Hk @⊥ @(absurd … Hk) @not_le_Sn_O | #h #t #Hind #k normalize cases k -k [ #Hk normalize @refl | #k #Hk normalize @(Hind k) @le_S_S_to_le @Hk ] ] qed. lemma nth_append_second: ∀A: Type[0].∀n.∀l1,l2:list A.∀d.n ≥ length A l1 -> nth n A (l1@l2) d = nth (n - length A l1) A l2 d. #A #n #l1 #l2 #d generalize in match n; -n; elim l1 [ normalize #k #Hk <(minus_n_O) @refl | #h #t #Hind #k normalize cases k -k; [ #Hk @⊥ @(absurd (S (|t|) ≤ 0)) [ @Hk | @not_le_Sn_O ] | #k #Hk normalize @(Hind k) @le_S_S_to_le @Hk ] ] qed. let rec fold_left_i_aux (A: Type[0]) (B: Type[0]) (f: nat → A → B → A) (x: A) (i: nat) (l: list B) on l ≝ match l with [ nil ⇒ x | cons hd tl ⇒ fold_left_i_aux A B f (f i x hd) (S i) tl ]. definition fold_left_i ≝ λA,B,f,x. fold_left_i_aux A B f x O. notation "hvbox(t⌈o ↦ h⌉)" with precedence 45 for @{ match (? : $o=$h) with [ refl ⇒ $t ] }. definition function_apply ≝ λA, B: Type[0]. λf: A → B. λa: A. f a. notation "f break $ x" left associative with precedence 99 for @{ 'function_apply $f $x }. interpretation "Function application" 'function_apply f x = (function_apply ? ? f x). let rec iterate (A: Type[0]) (f: A → A) (a: A) (n: nat) on n ≝ match n with [ O ⇒ a | S o ⇒ f (iterate A f a o) ]. let rec division_aux (m: nat) (n : nat) (p: nat) ≝ match ltb n (S p) with [ true ⇒ O | false ⇒ match m with [ O ⇒ O | (S q) ⇒ S (division_aux q (n - (S p)) p) ] ]. definition division ≝ λm, n: nat. match n with [ O ⇒ S m | S o ⇒ division_aux m m o ]. notation "hvbox(n break ÷ m)" right associative with precedence 47 for @{ 'division $n $m }. interpretation "Nat division" 'division n m = (division n m). let rec modulus_aux (m: nat) (n: nat) (p: nat) ≝ match leb n p with [ true ⇒ n | false ⇒ match m with [ O ⇒ n | S o ⇒ modulus_aux o (n - (S p)) p ] ]. definition modulus ≝ λm, n: nat. match n with [ O ⇒ m | S o ⇒ modulus_aux m m o ]. notation "hvbox(n break 'mod' m)" right associative with precedence 47 for @{ 'modulus $n $m }. interpretation "Nat modulus" 'modulus m n = (modulus m n). definition divide_with_remainder ≝ λm, n: nat. mk_Prod … (m ÷ n) (modulus m n). let rec exponential (m: nat) (n: nat) on n ≝ match n with [ O ⇒ S O | S o ⇒ m * exponential m o ]. interpretation "Nat exponential" 'exp n m = (exponential n m). notation "hvbox(a break ⊎ b)" left associative with precedence 55 for @{ 'disjoint_union $a $b }. interpretation "sum" 'disjoint_union A B = (Sum A B). theorem less_than_or_equal_monotone: ∀m, n: nat. m ≤ n → (S m) ≤ (S n). #m #n #H elim H /2 by le_n, le_S/ qed. theorem less_than_or_equal_b_complete: ∀m, n: nat. leb m n = false → ¬(m ≤ n). #m; elim m; normalize [ #n #H destruct | #y #H1 #z cases z normalize [ #H /2 by / | /3 by not_le_to_not_le_S_S/ ] ] qed. theorem less_than_or_equal_b_correct: ∀m, n: nat. leb m n = true → m ≤ n. #m elim m // #y #H1 #z cases z normalize [ #H destruct | #n #H lapply (H1 … H) /2 by le_S_S/ ] qed. definition less_than_or_equal_b_elim: ∀m, n: nat. ∀P: bool → Type[0]. (m ≤ n → P true) → (¬(m ≤ n) → P false) → P (leb m n). #m #n #P #H1 #H2; lapply (less_than_or_equal_b_correct m n) lapply (less_than_or_equal_b_complete m n) cases (leb m n) /3 by / qed. lemma lt_m_n_to_exists_o_plus_m_n: ∀m, n: nat. m < n → ∃o: nat. m + o = n. #m #n cases m [1: #irrelevant %{n} % |2: #m' #lt_hyp %{(n - S m')} >commutative_plus in ⊢ (??%?); refl % qed. lemma plus_left_monotone: ∀m, n, o: nat. m = n → o + m = o + n. #m #n #o #refl destruct % qed. lemma minus_plus_cancel: ∀m, n : nat. ∀proof: n ≤ m. (m - n) + n = m. #m #n #proof /2 by plus_minus/ qed. lemma lt_to_le_to_le: ∀n, m, p: nat. n < m → m ≤ p → n ≤ p. #n #m #p #H #H1 elim H [1: @(transitive_le n m p) /2/ |2: /2/ ] qed. lemma eqb_decidable: ∀l, r: nat. (eqb l r = true) ∨ (eqb l r = false). #l #r // qed. lemma r_Sr_and_l_r_to_Sl_r: ∀r, l: nat. (∃r': nat. r = S r' ∧ l = r') → S l = r. #r #l #exists_hyp cases exists_hyp #r' #and_hyp cases and_hyp #left_hyp #right_hyp destruct % qed. lemma eqb_Sn_to_exists_n': ∀m, n: nat. eqb (S m) n = true → ∃n': nat. n = S n'. #m #n cases n [1: normalize #absurd destruct(absurd) |2: #n' #_ %{n'} % ] qed. lemma eqb_true_to_eqb_S_S_true: ∀m, n: nat. eqb m n = true → eqb (S m) (S n) = true. #m #n normalize #assm assumption qed. lemma eqb_S_S_true_to_eqb_true: ∀m, n: nat. eqb (S m) (S n) = true → eqb m n = true. #m #n normalize #assm assumption qed. lemma eqb_true_to_refl: ∀l, r: nat. eqb l r = true → l = r. #l elim l [1: #r cases r [1: #_ % |2: #l' normalize #absurd destruct(absurd) ] |2: #l' #inductive_hypothesis #r #eqb_refl @r_Sr_and_l_r_to_Sl_r %{(pred r)} @conj [1: cases (eqb_Sn_to_exists_n' … eqb_refl) #r' #S_assm >S_assm % |2: cases (eqb_Sn_to_exists_n' … eqb_refl) #r' #refl_assm destruct normalize @inductive_hypothesis normalize in eqb_refl; assumption ] ] qed. lemma r_O_or_exists_r_r_Sr_and_l_neq_r_to_Sl_neq_r: ∀r, l: nat. r = O ∨ (∃r': nat. r = S r' ∧ l ≠ r') → S l ≠ r. #r #l #disj_hyp cases disj_hyp [1: #r_O_refl destruct @nmk #absurd destruct(absurd) |2: #exists_hyp cases exists_hyp #r' #conj_hyp cases conj_hyp #left_conj #right_conj destruct @nmk #S_S_refl_hyp elim right_conj #hyp @hyp // ] qed. lemma neq_l_r_to_neq_Sl_Sr: ∀l, r: nat. l ≠ r → S l ≠ S r. #l #r #l_neq_r_assm @nmk #Sl_Sr_assm cases l_neq_r_assm #assm @assm // qed. lemma eqb_false_to_not_refl: ∀l, r: nat. eqb l r = false → l ≠ r. #l elim l [1: #r cases r [1: normalize #absurd destruct(absurd) |2: #r' #_ @nmk #absurd destruct(absurd) ] |2: #l' #inductive_hypothesis #r cases r [1: #eqb_false_assm @r_O_or_exists_r_r_Sr_and_l_neq_r_to_Sl_neq_r @or_introl % |2: #r' #eqb_false_assm @neq_l_r_to_neq_Sl_Sr @inductive_hypothesis assumption ] ] qed. lemma le_to_lt_or_eq: ∀m, n: nat. m ≤ n → m = n ∨ m < n. #m #n #le_hyp cases le_hyp [1: @or_introl % |2: #m' #le_hyp' @or_intror normalize @le_S_S assumption ] qed. lemma le_neq_to_lt: ∀m, n: nat. m ≤ n → m ≠ n → m < n. #m #n #le_hyp #neq_hyp cases neq_hyp #eq_absurd_hyp generalize in match (le_to_lt_or_eq m n le_hyp); #disj_assm cases disj_assm [1: #absurd cases (eq_absurd_hyp absurd) |2: #assm assumption ] qed. inverter nat_jmdiscr for nat. lemma plus_lt_to_lt: ∀m, n, o: nat. m + n < o → m < o. #m #n #o elim n [1: <(plus_n_O m) in ⊢ (% → ?); #assumption assumption |2: #n' #inductive_hypothesis <(plus_n_Sm m n') in ⊢ (% → ?); #assm @inductive_hypothesis normalize in assm; normalize /2 by lt_S_to_lt/ ] qed. include "arithmetics/div_and_mod.ma". lemma n_plus_1_n_to_False: ∀n: nat. n + 1 = n → False. #n elim n [1: normalize #absurd destruct(absurd) |2: #n' #inductive_hypothesis normalize #absurd @inductive_hypothesis /2 by injective_S/ ] qed. lemma one_two_times_n_to_False: ∀n: nat. 1=2*n→False. #n cases n [1: normalize #absurd destruct(absurd) |2: #n' normalize #absurd lapply (injective_S … absurd) -absurd #absurd /2/ ] qed. let rec odd_p (n: nat) on n ≝ match n with [ O ⇒ False | S n' ⇒ even_p n' ] and even_p (n: nat) on n ≝ match n with [ O ⇒ True | S n' ⇒ odd_p n' ]. let rec n_even_p_to_n_plus_2_even_p (n: nat) on n: even_p n → even_p (n + 2) ≝ match n with [ O ⇒ ? | S n' ⇒ ? ] and n_odd_p_to_n_plus_2_odd_p (n: nat) on n: odd_p n → odd_p (n + 2) ≝ match n with [ O ⇒ ? | S n' ⇒ ? ]. [1,3: normalize #assm assumption |2: normalize @n_odd_p_to_n_plus_2_odd_p |4: normalize @n_even_p_to_n_plus_2_even_p ] qed. let rec two_times_n_even_p (n: nat) on n: even_p (2 * n) ≝ match n with [ O ⇒ ? | S n' ⇒ ? ] and two_times_n_plus_one_odd_p (n: nat) on n: odd_p ((2 * n) + 1) ≝ match n with [ O ⇒ ? | S n' ⇒ ? ]. [1,3: normalize @I |2: normalize >plus_n_Sm <(associative_plus n' n' 1) >(plus_n_O (n' + n')) cut(n' + n' + 0 + 1 = 2 * n' + 1) [1: // |2: #refl_assm >refl_assm @two_times_n_plus_one_odd_p ] |4: normalize >plus_n_Sm cut(n' + (n' + 1) + 1 = (2 * n') + 2) [1: normalize /2/ |2: #refl_assm >refl_assm @n_even_p_to_n_plus_2_even_p @two_times_n_even_p ] ] qed. include alias "basics/logic.ma". let rec even_p_to_not_odd_p (n: nat) on n: even_p n → ¬ odd_p n ≝ match n with [ O ⇒ ? | S n' ⇒ ? ] and odd_p_to_not_even_p (n: nat) on n: odd_p n → ¬ even_p n ≝ match n with [ O ⇒ ? | S n' ⇒ ? ]. [1: normalize #_ @nmk #assm assumption |3: normalize #absurd cases absurd |2: normalize @odd_p_to_not_even_p |4: normalize @even_p_to_not_odd_p ] qed. lemma even_p_odd_p_cases: ∀n: nat. even_p n ∨ odd_p n. #n elim n [1: normalize @or_introl @I |2: #n' #inductive_hypothesis normalize cases inductive_hypothesis #assm try (@or_introl assumption) try (@or_intror assumption) ] qed. lemma two_times_n_plus_one_refl_two_times_n_to_False: ∀m, n: nat. 2 * m + 1 = 2 * n → False. #m #n #assm cut (even_p (2 * n) ∧ even_p ((2 * m) + 1)) [1: >assm @conj @two_times_n_even_p |2: * #_ #absurd cases (even_p_to_not_odd_p … absurd) #assm @assm @two_times_n_plus_one_odd_p ] qed. lemma not_Some_neq_None_to_False: ∀a: Type[0]. ∀e: a. ¬ (Some a e ≠ None a) → False. #a #e #absurd cases absurd -absurd #absurd @absurd @nmk -absurd #absurd destruct(absurd) qed. lemma not_None_to_Some: ∀A: Type[0]. ∀o: option A. o ≠ None A → ∃v: A. o = Some A v. #A #o cases o [1: #absurd cases absurd #absurd' cases (absurd' (refl …)) |2: #v' #ignore /2/ ] qed. lemma inclusive_disjunction_true: ∀b, c: bool. (orb b c) = true → b = true ∨ c = true. # b # c elim b [ normalize # H @ or_introl % | normalize /3 by trans_eq, orb_true_l/ ] qed. lemma conjunction_true: ∀b, c: bool. andb b c = true → b = true ∧ c = true. # b # c elim b normalize [ /2 by conj/ | # K destruct ] qed. lemma eq_true_false: false=true → False. # K destruct qed. lemma inclusive_disjunction_b_true: ∀b. orb b true = true. # b cases b % qed. (* XXX: to be moved into logic/basics.ma *) lemma and_intro_right: ∀a, b: Prop. a → (a → b) → a ∧ b. #a #b /3/ qed. definition bool_to_Prop ≝ λb. match b with [ true ⇒ True | false ⇒ False ]. coercion bool_to_Prop: ∀b:bool. Prop ≝ bool_to_Prop on _b:bool to Type[0]. lemma bool_as_Prop_to_eq : ∀b : bool. b → b = true. **% qed. (* with this you can use prf : b with b : bool with rewriting >prf rewrites b as true *) coercion bool_to_Prop_to_eq : ∀b : bool.∀prf : b.b = true ≝ bool_as_Prop_to_eq on _prf : bool_to_Prop ? to (? = true). lemma andb_Prop : ∀b,d : bool.b → d → b∧d. #b #d #btrue #dtrue >btrue >dtrue % qed. lemma andb_Prop_true : ∀b,d : bool. (b∧d) → And (bool_to_Prop b) (bool_to_Prop d). #b #d #bdtrue elim (andb_true … bdtrue) #btrue #dtrue >btrue >dtrue % % qed. lemma orb_Prop_l : ∀b,d : bool.b → b∨d. #b #d #btrue >btrue % qed. lemma orb_Prop_r : ∀b,d : bool.d → b∨d. #b #d #dtrue >dtrue elim b % qed. lemma orb_Prop_true : ∀b,d : bool. (b∨d) → Or (bool_to_Prop b) (bool_to_Prop d). #b #d #bdtrue elim (orb_true_l … bdtrue) #xtrue >xtrue [%1 | %2] % qed. lemma notb_Prop : ∀b : bool. Not (bool_to_Prop b) → notb b. * * #H [@H % | %] qed. lemma eq_false_to_notb: ∀b. b = false → ¬ b. *; /2 by eq_true_false, I/ qed. lemma not_b_to_eq_false : ∀b : bool. Not (bool_to_Prop b) → b = false. ** #H [elim (H ?) % | %] qed. (* with this you can use prf : ¬b with b : bool with rewriting >prf rewrites b as false *) coercion not_bool_to_Prop_to_eq : ∀b : bool.∀prf : Not (bool_to_Prop b).b = false ≝ not_b_to_eq_false on _prf : Not (bool_to_Prop ?) to (? = false). lemma true_or_false_Prop : ∀b : bool.Or (bool_to_Prop b) (¬(bool_to_Prop b)). * [%1 % | %2 % *] qed. lemma eq_true_to_b : ∀b. b = true → b. #b #btrue >btrue % qed. definition if_then_else_safe : ∀A : Type[0].∀b : bool.(b → A) → (¬(bool_to_Prop b) → A) → A ≝ λA,b,f,g. match b return λx.match x with [true ⇒ bool_to_Prop b | false ⇒ ¬bool_to_Prop b] → A with [ true ⇒ f | false ⇒ g ] ?. elim b % * qed. notation > "'If' b 'then' 'with' ident prf1 'do' f 'else' 'with' ident prf2 'do' g" with precedence 46 for @{'if_then_else_safe $b (λ${ident prf1}.$f) (λ${ident prf2}.$g)}. notation > "'If' b 'then' 'with' ident prf1 : ty1 'do' f 'else' 'with' ident prf2 : ty2 'do' g" with precedence 46 for @{'if_then_else_safe $b (λ${ident prf1}:$ty1.$f) (λ${ident prf2}:$ty2.$g)}. notation > "'If' b 'then' f 'else' 'with' ident prf2 'do' g" with precedence 46 for @{'if_then_else_safe $b (λ_.$f) (λ${ident prf2}.$g)}. notation > "'If' b 'then' f 'else' 'with' ident prf2 : ty2 'do' g" with precedence 46 for @{'if_then_else_safe $b (λ_.$f) (λ${ident prf2}:$ty2.$g)}. notation > "'If' b 'then' 'with' ident prf1 'do' f 'else' g" with precedence 46 for @{'if_then_else_safe $b (λ${ident prf1}.$f) (λ_.$g)}. notation > "'If' b 'then' 'with' ident prf1 : ty1 'do' f 'else' g" with precedence 46 for @{'if_then_else_safe $b (λ${ident prf1}:$ty1.$f) (λ_.$g)}. notation < "hvbox('If' \nbsp b \nbsp 'then' \nbsp break 'with' \nbsp ident prf1 : ty1 \nbsp 'do' \nbsp break f \nbsp break 'else' \nbsp break 'with' \nbsp ident prf2 : ty2 \nbsp 'do' \nbsp break g)" with precedence 46 for @{'if_then_else_safe $b (λ${ident prf1}:$ty1.$f) (λ${ident prf2}:$ty2.$g)}. notation < "hvbox('If' \nbsp b \nbsp 'then' \nbsp break f \nbsp break 'else' \nbsp break 'with' \nbsp ident prf2 : ty2 \nbsp 'do' \nbsp break g)" with precedence 46 for @{'if_then_else_safe $b (λ_:$ty1.$f) (λ${ident prf2}:$ty2.$g)}. notation < "hvbox('If' \nbsp b \nbsp 'then' \nbsp break 'with' \nbsp ident prf1 : ty1 \nbsp 'do' \nbsp break f \nbsp break 'else' \nbsp break g)" with precedence 46 for @{'if_then_else_safe $b (λ${ident prf1}:$ty1.$f) (λ_:$ty2.$g)}. interpretation "dependent if then else" 'if_then_else_safe b f g = (if_then_else_safe ? b f g). (* Also extracts an equality proof (useful when not using Russell). *) notation > "hvbox('let' 〈ident x,ident y〉 'as' ident E 'return' ty ≝ t 'in' s)" with precedence 10 for @{ match $t return λx.∀${ident E}: x = $t. $ty with [ mk_Prod ${ident x} ${ident y} ⇒ λ${ident E}.$s ] (refl ? $t) }. notation > "hvbox('deplet' 〈ident x,ident y〉 'as' ident E ≝ t 'in' s)" with precedence 10 for @{ match $t return λx.∀${ident E}: x = $t. Σz: ?. ? with [ mk_Prod ${ident x} ${ident y} ⇒ λ${ident E}.$s ] (refl ? $t) }. notation > "hvbox('deplet' 〈ident x,ident y,ident z〉 'as' ident E ≝ t 'in' s)" with precedence 10 for @{ match $t return λx.∀${fresh w}:x = $t. Σq: ?. ? with [ mk_Prod ${fresh xy} ${ident z} ⇒ match ${fresh xy} return λx.∀${ident E}:? = $t. Σu: ?. ? with [ mk_Prod ${ident x} ${ident y} ⇒ λ${ident E}.$s ] ] (refl ? $t) }. notation > "hvbox('let' 〈ident x,ident y,ident z〉 'as' ident E 'return' ty ≝ t 'in' s)" with precedence 10 for @{ match $t return λx.∀${fresh w}:x = $t. Σq: ?. ? with [ mk_Prod ${fresh xy} ${ident z} ⇒ match ${fresh xy} return λx.∀${ident E}:? = $t. $ty with [ mk_Prod ${ident x} ${ident y} ⇒ λ${ident E}.$s ] ] (refl ? $t) }. lemma length_append: ∀A.∀l1,l2:list A. |l1 @ l2| = |l1| + |l2|. #A #l1 elim l1 [ // | #hd #tl #IH #l2 normalize Prop.∀s1:Σa1:A.P a1.∀s2:Σa2:A.P a2. s1 = s2 → (pi1 ?? s1) = (pi1 ?? s2). #A #P #s1 #s2 / by / qed. lemma not_neq_None_to_eq : ∀A.∀a : option A.¬a≠None ? → a = None ?. #A * [2: #a] * // #ABS elim (ABS ?) % #ABS' destruct(ABS') qed. coercion not_neq_None : ∀A.∀a : option A.∀prf : ¬a≠None ?.a = None ? ≝ not_neq_None_to_eq on _prf : ¬?≠None ? to ? = None ?. lemma None_Some_elim: ∀P:Prop. ∀A,a. None A = Some A a → P. #P #A #a #abs destruct qed. discriminator Prod. lemma pair_replace: ∀A,B,T:Type[0].∀P:T → Prop. ∀t. ∀a,b.∀c,c': A × B. c'=c → c ≃ 〈a,b〉 → P (t a b) → P (let 〈a',b'〉 ≝ c' in t a' b'). #A #B #T #P #t #a #b * #x #y * #x' #y' #H1 #H2 destruct // qed. lemma jmpair_as_replace: ∀A,B,T:Type[0].∀P:T → Prop. ∀a:A. ∀b:B.∀c,c': A × B. ∀t: ∀a,b. c' ≃ 〈a, b〉 → T. ∀EQc: c'= c. ∀EQ:c ≃ 〈a,b〉. P (t a b ?) → P ((let 〈a',b'〉 (*as H*) ≝ c' in λH. t a' b' H) (refl_jmeq ? c')). [2: >EQc @EQ |1: #A #B #T #P #a #b #c #c' #t #EQc >EQc in t; -c' normalize #f #EQ letin eta ≝ (? : ∀ab:A × B.∀P:A × B → Type[0]. P ab → P 〈\fst ab,\snd ab〉) [2: * // ] change with (eta 〈a,b〉 (λab.c ≃ ab) EQ) in match EQ; @(jmeq_elim ?? (λab.λH.P (f (\fst ab) (\snd ab) ?) → ?) ?? EQ) normalize -eta -EQ cases c in f; normalize // ] qed. lemma if_then_else_replace: ∀T: Type[0]. ∀P: T → Prop. ∀t1, t2: T. ∀c, c', c'': bool. c' = c → c ≃ c'' → P (if c'' then t1 else t2) → P (if c' then t1 else t2). #T #P #t1 #t2 #c #c' #c'' #c_refl #c_refl' destruct #assm assumption qed. lemma refl_to_jmrefl: ∀a: Type[0]. ∀l, r: a. l = r → l ≃ r. #a #l #r #refl destruct % qed. lemma prod_eq_left: ∀A: Type[0]. ∀p, q, r: A. p = q → 〈p, r〉 = 〈q, r〉. #A #p #q #r #hyp destruct % qed. lemma prod_eq_right: ∀A: Type[0]. ∀p, q, r: A. p = q → 〈r, p〉 = 〈r, q〉. #A #p #q #r #hyp destruct % qed. lemma destruct_bug_fix_1: ∀n: nat. S n = 0 → False. #n #absurd destruct(absurd) qed. lemma destruct_bug_fix_2: ∀m, n: nat. S m = S n → m = n. #m #n #refl destruct % qed. lemma eq_rect_Type1_r: ∀A: Type[1]. ∀a: A. ∀P: ∀x:A. eq ? x a → Type[1]. P a (refl A a) → ∀x: A.∀p:eq ? x a. P x p. #A #a #P #H #x #p generalize in match H; generalize in match P; cases p // qed. lemma Some_Some_elim: ∀T:Type[0].∀x,y:T.∀P:Type[2]. (x=y → P) → Some T x = Some T y → P. #T #x #y #P #H #K @H @option_destruct_Some // qed. lemma pair_destruct_right: ∀A: Type[0]. ∀B: Type[0]. ∀a, c: A. ∀b, d: B. 〈a, b〉 = 〈c, d〉 → b = d. #A #B #a #b #c #d // qed. lemma pose: ∀A:Type[0].∀B:A → Type[0].∀a:A. (∀a':A. a'=a → B a') → B a. /2/ qed. definition eq_sum: ∀A, B. (A → A → bool) → (B → B → bool) → (A ⊎ B) → (A ⊎ B) → bool ≝ λlt, rt, leq, req, left, right. match left with [ inl l ⇒ match right with [ inl l' ⇒ leq l l' | _ ⇒ false ] | inr r ⇒ match right with [ inr r' ⇒ req r r' | _ ⇒ false ] ]. lemma eq_sum_refl: ∀A, B: Type[0]. ∀leq, req. ∀s. ∀leq_refl: (∀t. leq t t = true). ∀req_refl: (∀u. req u u = true). eq_sum A B leq req s s = true. #A #B #leq #req #s #leq_refl #req_refl cases s assumption qed. definition eq_prod: ∀A, B. (A → A → bool) → (B → B → bool) → (A × B) → (A × B) → bool ≝ λlt, rt, leq, req, left, right. let 〈l, r〉 ≝ left in let 〈l', r'〉 ≝ right in leq l l' ∧ req r r'. lemma eq_prod_refl: ∀A, B: Type[0]. ∀leq, req. ∀s. ∀leq_refl: (∀t. leq t t = true). ∀req_refl: (∀u. req u u = true). eq_prod A B leq req s s = true. #A #B #leq #req #s #leq_refl #req_refl cases s whd in ⊢ (? → ? → ??%?); #l #r >leq_refl @req_refl qed. lemma geb_to_leb: ∀a,b:ℕ.geb a b = leb b a. #a #b / by refl/ qed. lemma nth_last: ∀A,a,l. nth (|l|) A l a = a. #A #a #l elim l [ / by / | #h #t #Hind whd in match (nth ????); whd in match (tail ??); @Hind ] qed. lemma minus_zero_to_le: ∀n,m:ℕ.n - m = 0 → n ≤ m. #n elim n [ #m #_ @le_O_n | #n' #Hind #m cases m [ #H -n whd in match (minus ??) in H; >H @le_n | #m' -m #H whd in match (minus ??) in H; @le_S_S @Hind @H ] ] qed. lemma plus_zero_zero: ∀n,m:ℕ.n + m = 0 → m = 0. #n #m #Hn @sym_eq @le_n_O_to_eq commutative_plus @le_plus_n_r qed. (* this can probably be done more simply somewhere *) lemma not_true_is_false: ∀b:bool.b ≠ true → b = false. #b cases b [ #H @⊥ @(absurd ?? H) @refl | #H @refl ] qed. (* this is in the stdlib, but commented out, why? *) theorem plus_Sn_m1: ∀n,m:nat. S m + n = m + S n. #n (elim n) normalize /2 by S_pred/ qed. (* these lemmas seem superfluous, but not sure how to replace them *) lemma double_plus_equal: ∀a,b:ℕ.a + a = b + b → a = b. #a elim a [ normalize #b // | -a #a #Hind #b cases b [ /2 by le_n_O_to_eq/ | -b #b normalize (Hind b (injective_S ?? (injective_S ?? H))) // ] ] qed. lemma sth_not_s: ∀x.x ≠ S x. #x cases x [ // | #y // ] qed. lemma le_to_eq_plus: ∀n,z. n ≤ z → ∃k.z = n + k. #n #z elim z [ #H cases (le_to_or_lt_eq … H) [ #H2 @⊥ @(absurd … H2) @not_le_Sn_O | #H2 @(ex_intro … 0) >H2 // ] | #z' #Hind #H cases (le_to_or_lt_eq … H) [ #H' elim (Hind (le_S_S_to_le … H')) #k' #H2 @(ex_intro … (S k')) >H2 >plus_n_Sm // | #H' @(ex_intro … 0) >H' // ] ] qed. lemma nth_safe_nth: ∀A:Type[0].∀l:list A.∀i.∀proof:i < |l|.∀x.nth_safe A i l proof = nth i A l x. #A #l elim l [ #i #prf @⊥ @(absurd ? prf) @le_to_not_lt @le_O_n | -l #h #t #Hind #i elim i [ #prf #x normalize @refl | -i #i #Hind2 #prf #x whd in match (nth ????); whd in match (tail ??); whd in match (nth_safe ????); @Hind ] ] qed. lemma flatten_singleton: ∀A,a. flatten A [a] = a. #A #a normalize // qed. lemma length_flatten_cons: ∀A,hd,tl. |flatten A (hd::tl)| = |hd| + |flatten A tl|. #A #hd #tl normalize // qed. lemma tech_transitive_lt_3: ∀n1,n2,n3,b. n1 < b → n2 < b → n3 < b → n1 + n2 + n3 < 3 * b. #n1 #n2 #n3 #b #H1 #H2 #H3 @(transitive_lt … (b + n2 + n3)) [ @monotonic_lt_plus_l @monotonic_lt_plus_l // ] @(transitive_lt … (b + b + n3)) [ @monotonic_lt_plus_l @monotonic_lt_plus_r // ] @(lt_to_le_to_lt … (b + b + b)) [ @monotonic_lt_plus_r // ] // qed. lemma m_lt_1_to_m_refl_0: ∀m: nat. m < 1 → m = 0. #m cases m try (#irrelevant %) #m' whd in ⊢ (% → ?); #relevant lapply (le_S_S_to_le … relevant) change with (? < ? → ?) -relevant #relevant cases (lt_to_not_zero … relevant) qed. lemma ltb_false_to_not_lt: ∀p,q:ℕ.ltb p q = false → p ≮ q. #p #q #H @leb_false_to_not_le @H qed. lemma ltb_true_to_lt: ∀p,q:ℕ.ltb p q = true → p < q. #p #q #H @leb_true_to_le @H qed. lemma plus_equals_zero: ∀x,y:ℕ.x + y = x → y = 0. #x #y cases y [ #_ @refl | -y #y elim x [