From: Wilmer Ricciotti Date: Fri, 6 Jul 2012 12:29:30 +0000 (+0000) Subject: manual commit after active hyperlinks X-Git-Tag: make_still_working~1621 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=c0ac63fead67ea1902e3d923ce877a2779cf501e manual commit after active hyperlinks --- diff --git a/weblib/arithmetics/nat.ma b/weblib/arithmetics/nat.ma index b735e1210..d45b0121b 100644 --- a/weblib/arithmetics/nat.ma +++ b/weblib/arithmetics/nat.ma @@ -11,7 +11,7 @@ include "basics/relations.ma". -inductive nat : Type[0] ≝ +img class="anchor" src="icons/tick.png" id="nat"inductive nat : Type[0] ≝ | O : nat | S : nat → nat. @@ -19,37 +19,37 @@ interpretation "Natural numbers" 'N = nat. alias num (instance 0) = "natural number". -definition pred ≝ +img class="anchor" src="icons/tick.png" id="pred"definition pred ≝ λn. match n with [ O ⇒ a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a | S p ⇒ p]. -theorem pred_Sn : ∀n.n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n). +img class="anchor" src="icons/tick.png" id="pred_Sn"theorem pred_Sn : ∀n.n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n). // qed. -theorem injective_S : a href="cic:/matita/basics/relations/injective.def(1)"injective/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a. +img class="anchor" src="icons/tick.png" id="injective_S"theorem injective_S : a href="cic:/matita/basics/relations/injective.def(1)"injective/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a. // qed. (* theorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m. //. qed. *) -theorem not_eq_S: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a m → a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m. +img class="anchor" src="icons/tick.png" id="not_eq_S"theorem not_eq_S: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a m → a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a/span/span/ qed. -definition not_zero: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop ≝ +img class="anchor" src="icons/tick.png" id="not_zero"definition not_zero: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop ≝ λn: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. match n with [ O ⇒ a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a | (S p) ⇒ a href="cic:/matita/basics/logic/True.ind(1,0,0)"True/a ]. -theorem not_eq_O_S : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n. +img class="anchor" src="icons/tick.png" id="not_eq_O_S"theorem not_eq_O_S : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n. #n @a href="cic:/matita/basics/logic/Not.con(0,1,1)"nmk/a #eqOS (change with (a href="cic:/matita/arithmetics/nat/not_zero.def(1)"not_zero/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a)) >eqOS // qed. -theorem not_eq_n_Sn: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n. +img class="anchor" src="icons/tick.png" id="not_eq_n_Sn"theorem not_eq_n_Sn: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n. #n (elim n) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/not_eq_S.def(4)"not_eq_S/a/span/span/ qed. -theorem nat_case: +img class="anchor" src="icons/tick.png" id="nat_case"theorem nat_case: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.∀P:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop. (na title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a → P a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a) → (∀m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. na title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m → P (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m)) → P n. #n #P (elim n) /span class="autotactic"2span class="autotrace" trace /span/span/ qed. -theorem nat_elim2 : +img class="anchor" src="icons/tick.png" id="nat_elim2"theorem nat_elim2 : ∀R:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop. (∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. R a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a n) → (∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. R (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n) a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a) @@ -57,18 +57,18 @@ theorem nat_elim2 : → ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. R n m. #R #ROn #RSO #RSS #n (elim n) // #n0 #Rn0m #m (cases m) /span class="autotactic"2span class="autotrace" trace /span/span/ qed. -theorem decidable_eq_nat : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/basics/logic/decidable.def(1)"decidable/a (na title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/am). +img class="anchor" src="icons/tick.png" id="decidable_eq_nat"theorem decidable_eq_nat : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/basics/logic/decidable.def(1)"decidable/a (na title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/am). @a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"nat_elim2/a #n [ (cases n) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ | /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a, a href="cic:/matita/basics/logic/sym_not_eq.def(4)"sym_not_eq/a/span/span/ | #m #Hind (cases Hind) /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a, a href="cic:/matita/arithmetics/nat/not_eq_S.def(4)"not_eq_S/a/span/span/] qed. (*************************** plus ******************************) -let rec plus n m ≝ +img class="anchor" src="icons/tick.png" id="plus"let rec plus n m ≝ match n with [ O ⇒ m | S p ⇒ a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a (plus p m) ]. interpretation "natural plus" 'plus x y = (plus x y). -theorem plus_O_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="natural plus" href="cic:/fakeuri.def(1)"+/an. +img class="anchor" src="icons/tick.png" id="plus_O_n"theorem plus_O_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="natural plus" href="cic:/fakeuri.def(1)"+/an. // qed. (* @@ -76,10 +76,10 @@ theorem plus_Sn_m: ∀n,m:nat. S (n + m) = S n + m. // qed. *) -theorem plus_n_O: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a na title="natural plus" href="cic:/fakeuri.def(1)"+/aa title="natural number" href="cic:/fakeuri.def(1)"0/a. +img class="anchor" src="icons/tick.png" id="plus_n_O"theorem plus_n_O: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a na title="natural plus" href="cic:/fakeuri.def(1)"+/aa title="natural number" href="cic:/fakeuri.def(1)"0/a. #n (elim n) normalize // qed. -theorem plus_n_Sm : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a (na title="natural plus" href="cic:/fakeuri.def(1)"+/am) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n a title="natural plus" href="cic:/fakeuri.def(1)"+/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m. +img class="anchor" src="icons/tick.png" id="plus_n_Sm"theorem plus_n_Sm : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a (na title="natural plus" href="cic:/fakeuri.def(1)"+/am) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n a title="natural plus" href="cic:/fakeuri.def(1)"+/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m. #n (elim n) normalize // qed. (* @@ -92,16 +92,16 @@ theorem plus_n_1 : ∀n:nat. S n = n+1. // qed. *) -theorem commutative_plus: a href="cic:/matita/basics/relations/commutative.def(1)"commutative/a ? a href="cic:/matita/arithmetics/nat/plus.fix(0,0,1)"plus/a. +img class="anchor" src="icons/tick.png" id="commutative_plus"theorem commutative_plus: a href="cic:/matita/basics/relations/commutative.def(1)"commutative/a ? a href="cic:/matita/arithmetics/nat/plus.fix(0,0,1)"plus/a. #n (elim n) normalize // qed. -theorem associative_plus : a href="cic:/matita/basics/relations/associative.def(1)"associative/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/plus.fix(0,0,1)"plus/a. +img class="anchor" src="icons/tick.png" id="associative_plus"theorem associative_plus : a href="cic:/matita/basics/relations/associative.def(1)"associative/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/plus.fix(0,0,1)"plus/a. #n (elim n) normalize // qed. -theorem assoc_plus1: ∀a,b,c. c a title="natural plus" href="cic:/fakeuri.def(1)"+/a (b a title="natural plus" href="cic:/fakeuri.def(1)"+/a a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b a title="natural plus" href="cic:/fakeuri.def(1)"+/a c a title="natural plus" href="cic:/fakeuri.def(1)"+/a a. +img class="anchor" src="icons/tick.png" id="assoc_plus1"theorem assoc_plus1: ∀a,b,c. c a title="natural plus" href="cic:/fakeuri.def(1)"+/a (b a title="natural plus" href="cic:/fakeuri.def(1)"+/a a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b a title="natural plus" href="cic:/fakeuri.def(1)"+/a c a title="natural plus" href="cic:/fakeuri.def(1)"+/a a. // qed. -theorem injective_plus_r: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/basics/relations/injective.def(1)"injective/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a (λm.na title="natural plus" href="cic:/fakeuri.def(1)"+/am). +img class="anchor" src="icons/tick.png" id="injective_plus_r"theorem injective_plus_r: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/basics/relations/injective.def(1)"injective/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a (λm.na title="natural plus" href="cic:/fakeuri.def(1)"+/am). #n (elim n) normalize /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/arithmetics/nat/injective_S.def(4)"injective_S/a/span/span/ qed. (* theorem inj_plus_r: \forall p,n,m:nat. p+n = p+m \to n=m @@ -115,43 +115,43 @@ theorem injective_plus_l: ∀m:nat.injective nat nat (λn.n+m). (*************************** times *****************************) -let rec times n m ≝ +img class="anchor" src="icons/tick.png" id="times"let rec times n m ≝ match n with [ O ⇒ a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a | S p ⇒ ma title="natural plus" href="cic:/fakeuri.def(1)"+/a(times p m) ]. interpretation "natural times" 'times x y = (times x y). -theorem times_Sn_m: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. ma title="natural plus" href="cic:/fakeuri.def(1)"+/ana title="natural times" href="cic:/fakeuri.def(1)"*/am a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a na title="natural times" href="cic:/fakeuri.def(1)"*/am. +img class="anchor" src="icons/tick.png" id="times_Sn_m"theorem times_Sn_m: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. ma title="natural plus" href="cic:/fakeuri.def(1)"+/ana title="natural times" href="cic:/fakeuri.def(1)"*/am a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a na title="natural times" href="cic:/fakeuri.def(1)"*/am. // qed. -theorem times_O_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="natural times" href="cic:/fakeuri.def(1)"*/an. +img class="anchor" src="icons/tick.png" id="times_O_n"theorem times_O_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="natural times" href="cic:/fakeuri.def(1)"*/an. // qed. -theorem times_n_O: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a na title="natural times" href="cic:/fakeuri.def(1)"*/aa href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a. +img class="anchor" src="icons/tick.png" id="times_n_O"theorem times_n_O: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a na title="natural times" href="cic:/fakeuri.def(1)"*/aa href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a. #n (elim n) // qed. -theorem times_n_Sm : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. na title="natural plus" href="cic:/fakeuri.def(1)"+/a(na title="natural times" href="cic:/fakeuri.def(1)"*/am) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a na title="natural times" href="cic:/fakeuri.def(1)"*/a(a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m). +img class="anchor" src="icons/tick.png" id="times_n_Sm"theorem times_n_Sm : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. na title="natural plus" href="cic:/fakeuri.def(1)"+/a(na title="natural times" href="cic:/fakeuri.def(1)"*/am) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a na title="natural times" href="cic:/fakeuri.def(1)"*/a(a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m). #n (elim n) normalize // qed. -theorem commutative_times : a href="cic:/matita/basics/relations/commutative.def(1)"commutative/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/times.fix(0,0,2)"times/a. +img class="anchor" src="icons/tick.png" id="commutative_times"theorem commutative_times : a href="cic:/matita/basics/relations/commutative.def(1)"commutative/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/times.fix(0,0,2)"times/a. #n (elim n) normalize // qed. (* variant sym_times : \forall n,m:nat. n*m = m*n \def symmetric_times. *) -theorem distributive_times_plus : a href="cic:/matita/basics/relations/distributive.def(1)"distributive/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/times.fix(0,0,2)"times/a a href="cic:/matita/arithmetics/nat/plus.fix(0,0,1)"plus/a. +img class="anchor" src="icons/tick.png" id="distributive_times_plus"theorem distributive_times_plus : a href="cic:/matita/basics/relations/distributive.def(1)"distributive/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/times.fix(0,0,2)"times/a a href="cic:/matita/arithmetics/nat/plus.fix(0,0,1)"plus/a. #n (elim n) normalize // qed. -theorem distributive_times_plus_r : +img class="anchor" src="icons/tick.png" id="distributive_times_plus_r"theorem distributive_times_plus_r : ∀a,b,c:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. (ba title="natural plus" href="cic:/fakeuri.def(1)"+/ac)a title="natural times" href="cic:/fakeuri.def(1)"*/aa a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a ba title="natural times" href="cic:/fakeuri.def(1)"*/aa a title="natural plus" href="cic:/fakeuri.def(1)"+/a ca title="natural times" href="cic:/fakeuri.def(1)"*/aa. // qed. -theorem associative_times: a href="cic:/matita/basics/relations/associative.def(1)"associative/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/times.fix(0,0,2)"times/a. +img class="anchor" src="icons/tick.png" id="associative_times"theorem associative_times: a href="cic:/matita/basics/relations/associative.def(1)"associative/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/times.fix(0,0,2)"times/a. #n (elim n) normalize // qed. -lemma times_times: ∀x,y,z. xa title="natural times" href="cic:/fakeuri.def(1)"*/a(ya title="natural times" href="cic:/fakeuri.def(1)"*/az) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a ya title="natural times" href="cic:/fakeuri.def(1)"*/a(xa title="natural times" href="cic:/fakeuri.def(1)"*/az). +img class="anchor" src="icons/tick.png" id="times_times"lemma times_times: ∀x,y,z. xa title="natural times" href="cic:/fakeuri.def(1)"*/a(ya title="natural times" href="cic:/fakeuri.def(1)"*/az) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a ya title="natural times" href="cic:/fakeuri.def(1)"*/a(xa title="natural times" href="cic:/fakeuri.def(1)"*/az). // qed. -theorem times_n_1 : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n a title="natural times" href="cic:/fakeuri.def(1)"*/a a title="natural number" href="cic:/fakeuri.def(1)"1/a. +img class="anchor" src="icons/tick.png" id="times_n_1"theorem times_n_1 : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n a title="natural times" href="cic:/fakeuri.def(1)"*/a a title="natural number" href="cic:/fakeuri.def(1)"1/a. #n // qed. (* ci servono questi risultati? @@ -182,7 +182,7 @@ qed. (******************** ordering relations ************************) -inductive le (n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a) : a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop ≝ +img class="anchor" src="icons/tick.png" id="le"inductive le (n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a) : a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop ≝ | le_n : le n n | le_S : ∀ m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. le n m → le n (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m). @@ -190,7 +190,7 @@ interpretation "natural 'less or equal to'" 'leq x y = (le x y). interpretation "natural 'neither less nor equal to'" 'nleq x y = (Not (le x y)). -definition lt: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop ≝ λn,m. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. +img class="anchor" src="icons/tick.png" id="lt"definition lt: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop ≝ λn,m. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. interpretation "natural 'less than'" 'lt x y = (lt x y). interpretation "natural 'not less than'" 'nless x y = (Not (lt x y)). @@ -198,16 +198,16 @@ interpretation "natural 'not less than'" 'nless x y = (Not (lt x y)). (* lemma eq_lt: ∀n,m. (n < m) = (S n ≤ m). // qed. *) -definition ge: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop ≝ λn,m.m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. +img class="anchor" src="icons/tick.png" id="ge"definition ge: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop ≝ λn,m.m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. interpretation "natural 'greater or equal to'" 'geq x y = (ge x y). -definition gt: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop ≝ λn,m.ma title="natural 'less than'" href="cic:/fakeuri.def(1)"</an. +img class="anchor" src="icons/tick.png" id="gt"definition gt: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop ≝ λn,m.ma title="natural 'less than'" href="cic:/fakeuri.def(1)"</an. interpretation "natural 'greater than'" 'gt x y = (gt x y). interpretation "natural 'not greater than'" 'ngtr x y = (Not (gt x y)). -theorem transitive_le : a href="cic:/matita/basics/relations/transitive.def(2)"transitive/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/le.ind(1,0,1)"le/a. +img class="anchor" src="icons/tick.png" id="transitive_le"theorem transitive_le : a href="cic:/matita/basics/relations/transitive.def(2)"transitive/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/le.ind(1,0,1)"le/a. #a #b #c #leab #lebc (elim lebc) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le.con(0,2,1)"le_S/a/span/span/ qed. @@ -215,29 +215,29 @@ qed. theorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p \def transitive_le. *) -theorem transitive_lt: a href="cic:/matita/basics/relations/transitive.def(2)"transitive/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/lt.def(1)"lt/a. +img class="anchor" src="icons/tick.png" id="transitive_lt"theorem transitive_lt: a href="cic:/matita/basics/relations/transitive.def(2)"transitive/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/lt.def(1)"lt/a. #a #b #c #ltab #ltbc (elim ltbc) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le.con(0,2,1)"le_S/a/span/span/qed. (* theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p \def transitive_lt. *) -theorem le_S_S: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m. +img class="anchor" src="icons/tick.png" id="le_S_S"theorem le_S_S: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m. #n #m #lenm (elim lenm) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a, a href="cic:/matita/arithmetics/nat/le.con(0,2,1)"le_S/a/span/span/ qed. -theorem le_O_n : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. +img class="anchor" src="icons/tick.png" id="le_O_n"theorem le_O_n : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. #n (elim n) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a, a href="cic:/matita/arithmetics/nat/le.con(0,2,1)"le_S/a/span/span/ qed. -theorem le_n_Sn : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n. +img class="anchor" src="icons/tick.png" id="le_n_Sn"theorem le_n_Sn : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a, a href="cic:/matita/arithmetics/nat/le.con(0,2,1)"le_S/a/span/span/ qed. -theorem le_pred_n : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. +img class="anchor" src="icons/tick.png" id="le_pred_n"theorem le_pred_n : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. #n (elim n) // qed. -theorem monotonic_pred: a href="cic:/matita/basics/relations/monotonic.def(1)"monotonic/a ? a href="cic:/matita/arithmetics/nat/le.ind(1,0,1)"le/a a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a. +img class="anchor" src="icons/tick.png" id="monotonic_pred"theorem monotonic_pred: a href="cic:/matita/basics/relations/monotonic.def(1)"monotonic/a ? a href="cic:/matita/arithmetics/nat/le.ind(1,0,1)"le/a a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a. #n #m #lenm (elim lenm) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a, a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a/span/span/ qed. -theorem le_S_S_to_le: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. +img class="anchor" src="icons/tick.png" id="le_S_S_to_le"theorem le_S_S_to_le: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. (* demo *) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_pred.def(4)"monotonic_pred/a/span/span/ qed. @@ -248,26 +248,26 @@ theorem lt_S_S_to_lt: ∀n,m. S n < S m → n < m. theorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m. /2/ qed. *) -theorem lt_to_not_zero : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → a href="cic:/matita/arithmetics/nat/not_zero.def(1)"not_zero/a m. +img class="anchor" src="icons/tick.png" id="lt_to_not_zero"theorem lt_to_not_zero : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → a href="cic:/matita/arithmetics/nat/not_zero.def(1)"not_zero/a m. #n #m #Hlt (elim Hlt) // qed. (* lt vs. le *) -theorem not_le_Sn_O: ∀ n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a. +img class="anchor" src="icons/tick.png" id="not_le_Sn_O"theorem not_le_Sn_O: ∀ n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a. #n @a href="cic:/matita/basics/logic/Not.con(0,1,1)"nmk/a #Hlen0 @(a href="cic:/matita/arithmetics/nat/lt_to_not_zero.def(2)"lt_to_not_zero/a ?? Hlen0) qed. -theorem not_le_to_not_le_S_S: ∀ n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a m → a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m. +img class="anchor" src="icons/tick.png" id="not_le_to_not_le_S_S"theorem not_le_to_not_le_S_S: ∀ n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a m → a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m. /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a, a href="cic:/matita/arithmetics/nat/monotonic_pred.def(4)"monotonic_pred/a/span/span/ qed. -theorem not_le_S_S_to_not_le: ∀ n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m → n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a m. +img class="anchor" src="icons/tick.png" id="not_le_S_S_to_not_le"theorem not_le_S_S_to_not_le: ∀ n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m → n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a m. /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a, a href="cic:/matita/arithmetics/nat/le_S_S.def(2)"le_S_S/a/span/span/ qed. -theorem decidable_le: ∀n,m. a href="cic:/matita/basics/logic/decidable.def(1)"decidable/a (na title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/am). +img class="anchor" src="icons/tick.png" id="decidable_le"theorem decidable_le: ∀n,m. a href="cic:/matita/basics/logic/decidable.def(1)"decidable/a (na title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/am). @a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"nat_elim2/a #n /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ #m * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a, a href="cic:/matita/arithmetics/nat/not_le_to_not_le_S_S.def(5)"not_le_to_not_le_S_S/a, a href="cic:/matita/arithmetics/nat/le_S_S.def(2)"le_S_S/a/span/span/ qed. -theorem decidable_lt: ∀n,m. a href="cic:/matita/basics/logic/decidable.def(1)"decidable/a (n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m). +img class="anchor" src="icons/tick.png" id="decidable_lt"theorem decidable_lt: ∀n,m. a href="cic:/matita/basics/logic/decidable.def(1)"decidable/a (n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m). #n #m @a href="cic:/matita/arithmetics/nat/decidable_le.def(6)"decidable_le/a qed. -theorem not_le_Sn_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a n. +img class="anchor" src="icons/tick.png" id="not_le_Sn_n"theorem not_le_Sn_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a n. #n (elim n) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/not_le_to_not_le_S_S.def(5)"not_le_to_not_le_S_S/a/span/span/ qed. (* this is le_S_S_to_le @@ -275,10 +275,10 @@ theorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m. /2/ qed. *) -lemma le_gen: ∀P:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop.∀n.(∀i. i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → P i) → P n. +img class="anchor" src="icons/tick.png" id="le_gen"lemma le_gen: ∀P:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop.∀n.(∀i. i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → P i) → P n. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a/span/span/ qed. -theorem not_le_to_lt: ∀n,m. n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a m → m a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n. +img class="anchor" src="icons/tick.png" id="not_le_to_lt"theorem not_le_to_lt: ∀n,m. n a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a m → m a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n. @a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"nat_elim2/a #n [#abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ |/span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_S_S.def(2)"le_S_S/a/span/span/ @@ -286,27 +286,27 @@ theorem not_le_to_lt: ∀n,m. n a title="natural 'neither less nor equal to'" h ] qed. -theorem lt_to_not_le: ∀n,m. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → m a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a n. +img class="anchor" src="icons/tick.png" id="lt_to_not_le"theorem lt_to_not_le: ∀n,m. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → m a title="natural 'neither less nor equal to'" href="cic:/fakeuri.def(1)"≰/a n. #n #m #Hltnm (elim Hltnm) /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a, a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a/span/span/ qed. -theorem not_lt_to_le: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'not less than'" href="cic:/fakeuri.def(1)"≮/a m → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. +img class="anchor" src="icons/tick.png" id="not_lt_to_le"theorem not_lt_to_le: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'not less than'" href="cic:/fakeuri.def(1)"≮/a m → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. /span class="autotactic"4span class="autotrace" trace a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a, a href="cic:/matita/arithmetics/nat/not_le_to_lt.def(5)"not_le_to_lt/a, a href="cic:/matita/arithmetics/nat/monotonic_pred.def(4)"monotonic_pred/a/span/span/ qed. -theorem le_to_not_lt: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → m a title="natural 'not less than'" href="cic:/fakeuri.def(1)"≮/a n. +img class="anchor" src="icons/tick.png" id="le_to_not_lt"theorem le_to_not_lt: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → m a title="natural 'not less than'" href="cic:/fakeuri.def(1)"≮/a n. #n #m #H @a href="cic:/matita/arithmetics/nat/lt_to_not_le.def(7)"lt_to_not_le/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_S_S.def(2)"le_S_S/a/span/span/ (* /3/ *) qed. (* lt and le trans *) -theorem lt_to_le_to_lt: ∀n,m,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a p → n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a p. +img class="anchor" src="icons/tick.png" id="lt_to_le_to_lt"theorem lt_to_le_to_lt: ∀n,m,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a p → n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a p. #n #m #p #H #H1 (elim H1) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/transitive_lt.def(3)"transitive_lt/a, a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a/span/span/ qed. -theorem le_to_lt_to_lt: ∀n,m,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → m a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a p → n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a p. +img class="anchor" src="icons/tick.png" id="le_to_lt_to_lt"theorem le_to_lt_to_lt: ∀n,m,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → m a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a p → n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a p. #n #m #p #H (elim H) /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/arithmetics/nat/transitive_lt.def(3)"transitive_lt/a, a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a/span/span/ qed. -theorem lt_S_to_lt: ∀n,m. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m. +img class="anchor" src="icons/tick.png" id="lt_S_to_lt"theorem lt_S_to_lt: ∀n,m. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/transitive_lt.def(3)"transitive_lt/a, a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a/span/span/ qed. -theorem ltn_to_ltO: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m. +img class="anchor" src="icons/tick.png" id="ltn_to_ltO"theorem ltn_to_ltO: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_to_lt_to_lt.def(4)"le_to_lt_to_lt/a/span/span/ qed. (* @@ -320,12 +320,12 @@ apply (ltn_to_ltO (pred (S O)) (pred n) ?). ] qed. *) -theorem lt_O_n_elim: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n → +img class="anchor" src="icons/tick.png" id="lt_O_n_elim"theorem lt_O_n_elim: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n → ∀P:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop.(∀m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.P (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m)) → P n. #n (elim n) // #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ qed. -theorem S_pred: ∀n. a title="natural number" href="cic:/fakeuri.def(1)"0/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n → a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a(a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a n) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n. +img class="anchor" src="icons/tick.png" id="S_pred"theorem S_pred: ∀n. a title="natural number" href="cic:/fakeuri.def(1)"0/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n → a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a(a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a n) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n. #n #posn (cases posn) // qed. @@ -356,11 +356,11 @@ qed. *) (* le to lt or eq *) -theorem le_to_or_lt_eq: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m a title="logical or" href="cic:/fakeuri.def(1)"∨/a n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a m. +img class="anchor" src="icons/tick.png" id="le_to_or_lt_eq"theorem le_to_or_lt_eq: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m a title="logical or" href="cic:/fakeuri.def(1)"∨/a n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a m. #n #m #lenm (elim lenm) /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a, a href="cic:/matita/arithmetics/nat/le_to_lt_to_lt.def(4)"le_to_lt_to_lt/a, a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a/span/span/ qed. (* not eq *) -theorem lt_to_not_eq : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → n a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a m. +img class="anchor" src="icons/tick.png" id="lt_to_not_eq"theorem lt_to_not_eq : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → n a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a m. #n #m #H @a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a, a href="cic:/matita/basics/logic/Not.con(0,1,1)"nmk/a/span/span/ qed. (*not lt @@ -388,30 +388,30 @@ intro apply (not_le_Sn_n ? H3). qed. *) -theorem not_eq_to_le_to_lt: ∀n,m. na title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/am → na title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/am → na title="natural 'less than'" href="cic:/fakeuri.def(1)"</am. +img class="anchor" src="icons/tick.png" id="not_eq_to_le_to_lt"theorem not_eq_to_le_to_lt: ∀n,m. na title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/am → na title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/am → na title="natural 'less than'" href="cic:/fakeuri.def(1)"</am. #n #m #Hneq #Hle cases (a href="cic:/matita/arithmetics/nat/le_to_or_lt_eq.def(5)"le_to_or_lt_eq/a ?? Hle) // #Heq /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a, a href="cic:/matita/arithmetics/nat/not_le_to_lt.def(5)"not_le_to_lt/a/span/span/ qed. (* nelim (Hneq Heq) qed. *) (* le elimination *) -theorem le_n_O_to_eq : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a → a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/an. +img class="anchor" src="icons/tick.png" id="le_n_O_to_eq"theorem le_n_O_to_eq : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a → a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/an. #n (cases n) // #a #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ qed. -theorem le_n_O_elim: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a → ∀P: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a →Prop. P a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a → P n. +img class="anchor" src="icons/tick.png" id="le_n_O_elim"theorem le_n_O_elim: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a → ∀P: a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a →Prop. P a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a → P n. #n (cases n) // #a #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ qed. -theorem le_n_Sm_elim : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m → +img class="anchor" src="icons/tick.png" id="le_n_Sm_elim"theorem le_n_Sm_elim : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m → ∀P:Prop. (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m → P) → (na title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m → P) → P. #n #m #Hle #P (elim Hle) /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_S_S.def(2)"le_S_S/a/span/span/ qed. (* le and eq *) -theorem le_to_le_to_eq: ∀n,m. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a m. +img class="anchor" src="icons/tick.png" id="le_to_le_to_eq"theorem le_to_le_to_eq: ∀n,m. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a m. @a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"nat_elim2/a /span class="autotactic"4span class="autotrace" trace a href="cic:/matita/basics/logic/eq_f.def(3)"eq_f/a, a href="cic:/matita/arithmetics/nat/le_n_O_to_eq.def(4)"le_n_O_to_eq/a, a href="cic:/matita/arithmetics/nat/monotonic_pred.def(4)"monotonic_pred/a/span/span/ qed. -theorem lt_O_S : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n. +img class="anchor" src="icons/tick.png" id="lt_O_S"theorem lt_O_S : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/not_le_to_lt.def(5)"not_le_to_lt/a/span/span/ qed. (* @@ -441,7 +441,7 @@ qed. (* well founded induction principles *) -theorem nat_elim1 : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.∀P:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop. +img class="anchor" src="icons/tick.png" id="nat_elim1"theorem nat_elim1 : ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.∀P:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → Prop. (∀m.(∀p. p a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → P p) → P m) → P n. #n #P #H cut (∀q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. q a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → P q) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a/span/span/ @@ -456,25 +456,25 @@ qed. (* some properties of functions *) -definition increasing ≝ λf:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. f n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a f (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n). +img class="anchor" src="icons/tick.png" id="increasing"definition increasing ≝ λf:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. f n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a f (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n). -theorem increasing_to_monotonic: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. +img class="anchor" src="icons/tick.png" id="increasing_to_monotonic"theorem increasing_to_monotonic: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/increasing.def(2)"increasing/a f → a href="cic:/matita/basics/relations/monotonic.def(1)"monotonic/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/lt.def(1)"lt/a f. #f #incr #n #m #ltnm (elim ltnm) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/transitive_lt.def(3)"transitive_lt/a/span/span/ qed. -theorem le_n_fn: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. +img class="anchor" src="icons/tick.png" id="le_n_fn"theorem le_n_fn: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/increasing.def(2)"increasing/a f → ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a f n. #f #incr #n (elim n) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_to_lt_to_lt.def(4)"le_to_lt_to_lt/a/span/span/ qed. -theorem increasing_to_le: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. +img class="anchor" src="icons/tick.png" id="increasing_to_le"theorem increasing_to_le: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/increasing.def(2)"increasing/a f → ∀m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a title="exists" href="cic:/fakeuri.def(1)"∃/ai.m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a f i. #f #incr #m (elim m) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a/span/span/#n * #a #lenfa @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a ?? (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a a)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_to_lt_to_lt.def(4)"le_to_lt_to_lt/a/span/span/ qed. -theorem increasing_to_le2: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/increasing.def(2)"increasing/a f → +img class="anchor" src="icons/tick.png" id="increasing_to_le2"theorem increasing_to_le2: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/increasing.def(2)"increasing/a f → ∀m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. f a title="natural number" href="cic:/fakeuri.def(1)"0/a a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → a title="exists" href="cic:/fakeuri.def(1)"∃/ai. f i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m a title="logical and" href="cic:/fakeuri.def(1)"∧/a m a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a f (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a i). #f #incr #m #lem (elim lem) [@(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a ? ? a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a, a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a/span/span/ @@ -485,7 +485,7 @@ theorem increasing_to_le2: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0 ] qed. -theorem increasing_to_injective: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. +img class="anchor" src="icons/tick.png" id="increasing_to_injective"theorem increasing_to_injective: ∀f:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a → a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/increasing.def(2)"increasing/a f → a href="cic:/matita/basics/relations/injective.def(1)"injective/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a f. #f #incr #n #m cases(a href="cic:/matita/arithmetics/nat/decidable_le.def(6)"decidable_le/a n m) [#lenm cases(a href="cic:/matita/arithmetics/nat/le_to_or_lt_eq.def(5)"le_to_or_lt_eq/a … lenm) // @@ -497,7 +497,7 @@ theorem increasing_to_injective: ∀f:a href="cic:/matita/arithmetics/nat/nat.i qed. (*********************** monotonicity ***************************) -theorem monotonic_le_plus_r: +img class="anchor" src="icons/tick.png" id="monotonic_le_plus_r"theorem monotonic_le_plus_r: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/basics/relations/monotonic.def(1)"monotonic/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/le.ind(1,0,1)"le/a (λm.n a title="natural plus" href="cic:/fakeuri.def(1)"+/a m). #n #a #b (elim n) normalize // #m #H #leab @a href="cic:/matita/arithmetics/nat/le_S_S.def(2)"le_S_S/a /span class="autotactic"2span class="autotrace" trace /span/span/ qed. @@ -506,7 +506,7 @@ theorem monotonic_le_plus_r: theorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m ≝ monotonic_le_plus_r. *) -theorem monotonic_le_plus_l: +img class="anchor" src="icons/tick.png" id="monotonic_le_plus_l"theorem monotonic_le_plus_l: ∀m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/basics/relations/monotonic.def(1)"monotonic/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/le.ind(1,0,1)"le/a (λn.n a title="natural plus" href="cic:/fakeuri.def(1)"+/a m). /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_plus_r.def(3)"monotonic_le_plus_r/a/span/span/ qed. @@ -514,35 +514,35 @@ theorem monotonic_le_plus_l: theorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p \def monotonic_le_plus_l. *) -theorem le_plus: ∀n1,n2,m1,m2:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n1 a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n2 → m1 a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m2 +img class="anchor" src="icons/tick.png" id="le_plus"theorem le_plus: ∀n1,n2,m1,m2:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n1 a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n2 → m1 a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m2 → n1 a title="natural plus" href="cic:/fakeuri.def(1)"+/a m1 a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n2 a title="natural plus" href="cic:/fakeuri.def(1)"+/a m2. #n1 #n2 #m1 #m2 #len #lem @(a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a ? (n1a title="natural plus" href="cic:/fakeuri.def(1)"+/am2)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_plus_l.def(6)"monotonic_le_plus_l/a, a href="cic:/matita/arithmetics/nat/monotonic_le_plus_r.def(3)"monotonic_le_plus_r/a/span/span/ qed. -theorem le_plus_n :∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n a title="natural plus" href="cic:/fakeuri.def(1)"+/a m. +img class="anchor" src="icons/tick.png" id="le_plus_n"theorem le_plus_n :∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n a title="natural plus" href="cic:/fakeuri.def(1)"+/a m. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_plus_l.def(6)"monotonic_le_plus_l/a/span/span/ qed. -lemma le_plus_a: ∀a,n,m. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a a title="natural plus" href="cic:/fakeuri.def(1)"+/a m. +img class="anchor" src="icons/tick.png" id="le_plus_a"lemma le_plus_a: ∀a,n,m. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a a title="natural plus" href="cic:/fakeuri.def(1)"+/a m. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_plus.def(7)"le_plus/a/span/span/ qed. -lemma le_plus_b: ∀b,n,m. n a title="natural plus" href="cic:/fakeuri.def(1)"+/a b a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. +img class="anchor" src="icons/tick.png" id="le_plus_b"lemma le_plus_b: ∀b,n,m. n a title="natural plus" href="cic:/fakeuri.def(1)"+/a b a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a/span/span/ qed. -theorem le_plus_n_r :∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m a title="natural plus" href="cic:/fakeuri.def(1)"+/a n. +img class="anchor" src="icons/tick.png" id="le_plus_n_r"theorem le_plus_n_r :∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m a title="natural plus" href="cic:/fakeuri.def(1)"+/a n. /span class="autotactic"2span class="autotrace" trace /span/span/ qed. -theorem eq_plus_to_le: ∀n,m,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.na title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ama title="natural plus" href="cic:/fakeuri.def(1)"+/ap → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. +img class="anchor" src="icons/tick.png" id="eq_plus_to_le"theorem eq_plus_to_le: ∀n,m,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.na title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ama title="natural plus" href="cic:/fakeuri.def(1)"+/ap → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. // qed. -theorem le_plus_to_le: ∀a,n,m. a a title="natural plus" href="cic:/fakeuri.def(1)"+/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a a title="natural plus" href="cic:/fakeuri.def(1)"+/a m → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. +img class="anchor" src="icons/tick.png" id="le_plus_to_le"theorem le_plus_to_le: ∀a,n,m. a a title="natural plus" href="cic:/fakeuri.def(1)"+/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a a title="natural plus" href="cic:/fakeuri.def(1)"+/a m → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. #a (elim a) normalize /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_pred.def(4)"monotonic_pred/a/span/span/ qed. -theorem le_plus_to_le_r: ∀a,n,m. n a title="natural plus" href="cic:/fakeuri.def(1)"+/a a a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m a title="natural plus" href="cic:/fakeuri.def(1)"+/aa → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. +img class="anchor" src="icons/tick.png" id="le_plus_to_le_r"theorem le_plus_to_le_r: ∀a,n,m. n a title="natural plus" href="cic:/fakeuri.def(1)"+/a a a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m a title="natural plus" href="cic:/fakeuri.def(1)"+/aa → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_plus_to_le.def(5)"le_plus_to_le/a/span/span/ qed. (* plus & lt *) -theorem monotonic_lt_plus_r: +img class="anchor" src="icons/tick.png" id="monotonic_lt_plus_r"theorem monotonic_lt_plus_r: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/basics/relations/monotonic.def(1)"monotonic/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/lt.def(1)"lt/a (λm.na title="natural plus" href="cic:/fakeuri.def(1)"+/am). /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/increasing_to_monotonic.def(4)"increasing_to_monotonic/a/span/span/ qed. @@ -550,7 +550,7 @@ theorem monotonic_lt_plus_r: variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def monotonic_lt_plus_r. *) -theorem monotonic_lt_plus_l: +img class="anchor" src="icons/tick.png" id="monotonic_lt_plus_l"theorem monotonic_lt_plus_l: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/basics/relations/monotonic.def(1)"monotonic/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/lt.def(1)"lt/a (λm.ma title="natural plus" href="cic:/fakeuri.def(1)"+/an). (* /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/increasing_to_monotonic.def(4)"increasing_to_monotonic/a/span/span/ *) #n @a href="cic:/matita/arithmetics/nat/increasing_to_monotonic.def(4)"increasing_to_monotonic/a // qed. @@ -558,14 +558,14 @@ theorem monotonic_lt_plus_l: variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def monotonic_lt_plus_l. *) -theorem lt_plus: ∀n,m,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → p a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a q → n a title="natural plus" href="cic:/fakeuri.def(1)"+/a p a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m a title="natural plus" href="cic:/fakeuri.def(1)"+/a q. +img class="anchor" src="icons/tick.png" id="lt_plus"theorem lt_plus: ∀n,m,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → p a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a q → n a title="natural plus" href="cic:/fakeuri.def(1)"+/a p a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m a title="natural plus" href="cic:/fakeuri.def(1)"+/a q. #n #m #p #q #ltnm #ltpq -@(a href="cic:/matita/arithmetics/nat/transitive_lt.def(3)"transitive_lt/a ? (na title="natural plus" href="cic:/fakeuri.def(1)"+/aq))/span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_plus_r.def(3)"monotonic_le_plus_r/a, a href="cic:/matita/arithmetics/nat/monotonic_lt_plus_l.def(9)"monotonic_lt_plus_l/a/span/span/ qed. +@(a href="cic:/matita/arithmetics/nat/transitive_lt.def(3)"transitive_lt/a ? (na title="natural plus" href="cic:/fakeuri.def(1)"+/aq)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_lt_plus_l.def(5)"monotonic_lt_plus_l/a, a href="cic:/matita/arithmetics/nat/monotonic_le_plus_r.def(3)"monotonic_le_plus_r/a/span/span/ qed. -theorem lt_plus_to_lt_l :∀n,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. pa title="natural plus" href="cic:/fakeuri.def(1)"+/an a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a qa title="natural plus" href="cic:/fakeuri.def(1)"+/an → pa title="natural 'less than'" href="cic:/fakeuri.def(1)"</aq. +img class="anchor" src="icons/tick.png" id="lt_plus_to_lt_l"theorem lt_plus_to_lt_l :∀n,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. pa title="natural plus" href="cic:/fakeuri.def(1)"+/an a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a qa title="natural plus" href="cic:/fakeuri.def(1)"+/an → pa title="natural 'less than'" href="cic:/fakeuri.def(1)"</aq. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_plus_to_le.def(5)"le_plus_to_le/a/span/span/ qed. -theorem lt_plus_to_lt_r :∀n,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. na title="natural plus" href="cic:/fakeuri.def(1)"+/ap a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a na title="natural plus" href="cic:/fakeuri.def(1)"+/aq → pa title="natural 'less than'" href="cic:/fakeuri.def(1)"</aq. +img class="anchor" src="icons/tick.png" id="lt_plus_to_lt_r"theorem lt_plus_to_lt_r :∀n,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. na title="natural plus" href="cic:/fakeuri.def(1)"+/ap a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a na title="natural plus" href="cic:/fakeuri.def(1)"+/aq → pa title="natural 'less than'" href="cic:/fakeuri.def(1)"</aq. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/lt_plus_to_lt_l.def(6)"lt_plus_to_lt_l/a/span/span/ qed. (* @@ -577,7 +577,7 @@ normalize napplyS le_plus // qed. *) (* times *) -theorem monotonic_le_times_r: +img class="anchor" src="icons/tick.png" id="monotonic_le_times_r"theorem monotonic_le_times_r: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/basics/relations/monotonic.def(1)"monotonic/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/le.ind(1,0,1)"le/a (λm. n a title="natural times" href="cic:/fakeuri.def(1)"*/a m). #n #x #y #lexy (elim n) normalize//(* lento /2/*) #a #lea @a href="cic:/matita/arithmetics/nat/le_plus.def(7)"le_plus/a // @@ -597,16 +597,16 @@ theorem monotonic_le_times_l: theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p \def monotonic_le_times_l. *) -theorem le_times: ∀n1,n2,m1,m2:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. +img class="anchor" src="icons/tick.png" id="le_times"theorem le_times: ∀n1,n2,m1,m2:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n1 a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n2 → m1 a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m2 → n1a title="natural times" href="cic:/fakeuri.def(1)"*/am1 a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n2a title="natural times" href="cic:/fakeuri.def(1)"*/am2. #n1 #n2 #m1 #m2 #len #lem @(a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a ? (n1a title="natural times" href="cic:/fakeuri.def(1)"*/am2)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_times_r.def(8)"monotonic_le_times_r/a/span/span/ qed. (* interessante *) -theorem lt_times_n: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a na title="natural times" href="cic:/fakeuri.def(1)"*/am. +img class="anchor" src="icons/tick.png" id="lt_times_n"theorem lt_times_n: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a na title="natural times" href="cic:/fakeuri.def(1)"*/am. #n #m #H /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_times_r.def(8)"monotonic_le_times_r/a/span/span/ qed. -theorem le_times_to_le: +img class="anchor" src="icons/tick.png" id="le_times_to_le"theorem le_times_to_le: ∀a,n,m. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a a → a a title="natural times" href="cic:/fakeuri.def(1)"*/a n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a a title="natural times" href="cic:/fakeuri.def(1)"*/a m → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. #a @a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"nat_elim2/a normalize [// @@ -683,7 +683,7 @@ simplify.rewrite < plus_n_O.rewrite < plus_n_O.assumption. apply lt_plus.assumption.assumption. qed. *) -theorem monotonic_lt_times_r: +img class="anchor" src="icons/tick.png" id="monotonic_lt_times_r"theorem monotonic_lt_times_r: ∀c:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a c → a href="cic:/matita/basics/relations/monotonic.def(1)"monotonic/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/lt.def(1)"lt/a (λt.(ca title="natural times" href="cic:/fakeuri.def(1)"*/at)). #c #posc #n #m #ltnm (elim ltnm) normalize @@ -692,12 +692,12 @@ theorem monotonic_lt_times_r: ] qed. -theorem monotonic_lt_times_l: +img class="anchor" src="icons/tick.png" id="monotonic_lt_times_l"theorem monotonic_lt_times_l: ∀c:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a c → a href="cic:/matita/basics/relations/monotonic.def(1)"monotonic/a a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a a href="cic:/matita/arithmetics/nat/lt.def(1)"lt/a (λt.(ta title="natural times" href="cic:/fakeuri.def(1)"*/ac)). /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_lt_times_r.def(9)"monotonic_lt_times_r/a/span/span/ qed. -theorem lt_to_le_to_lt_times: +img class="anchor" src="icons/tick.png" id="lt_to_le_to_lt_times"theorem lt_to_le_to_lt_times: ∀n,m,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → p a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a q → a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a q → na title="natural times" href="cic:/fakeuri.def(1)"*/ap a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a ma title="natural times" href="cic:/fakeuri.def(1)"*/aq. #n #m #p #q #ltnm #lepq #posq @(a href="cic:/matita/arithmetics/nat/le_to_lt_to_lt.def(4)"le_to_lt_to_lt/a ? (na title="natural times" href="cic:/fakeuri.def(1)"*/aq)) @@ -706,18 +706,18 @@ theorem lt_to_le_to_lt_times: ] qed. -theorem lt_times:∀n,m,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. na title="natural 'less than'" href="cic:/fakeuri.def(1)"</am → pa title="natural 'less than'" href="cic:/fakeuri.def(1)"</aq → na title="natural times" href="cic:/fakeuri.def(1)"*/ap a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a ma title="natural times" href="cic:/fakeuri.def(1)"*/aq. -#n #m #p #q #ltnm #ltpq @a href="cic:/matita/arithmetics/nat/lt_to_le_to_lt_times.def(12)"lt_to_le_to_lt_times/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_plus_b.def(8)"le_plus_b/a, a href="cic:/matita/arithmetics/nat/ltn_to_ltO.def(5)"ltn_to_ltO/a/span/span/ +img class="anchor" src="icons/tick.png" id="lt_times"theorem lt_times:∀n,m,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. na title="natural 'less than'" href="cic:/fakeuri.def(1)"</am → pa title="natural 'less than'" href="cic:/fakeuri.def(1)"</aq → na title="natural times" href="cic:/fakeuri.def(1)"*/ap a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a ma title="natural times" href="cic:/fakeuri.def(1)"*/aq. +#n #m #p #q #ltnm #ltpq @a href="cic:/matita/arithmetics/nat/lt_to_le_to_lt_times.def(11)"lt_to_le_to_lt_times/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_plus_b.def(8)"le_plus_b/a, a href="cic:/matita/arithmetics/nat/ltn_to_ltO.def(5)"ltn_to_ltO/a/span/span/ qed. -theorem lt_times_n_to_lt_l: +img class="anchor" src="icons/tick.png" id="lt_times_n_to_lt_l"theorem lt_times_n_to_lt_l: ∀n,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. pa title="natural times" href="cic:/fakeuri.def(1)"*/an a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a qa title="natural times" href="cic:/fakeuri.def(1)"*/an → p a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a q. #n #p #q #Hlt (elim (a href="cic:/matita/arithmetics/nat/decidable_lt.def(7)"decidable_lt/a p q)) // #nltpq @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a @(a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a ? ? (a href="cic:/matita/arithmetics/nat/lt_to_not_le.def(7)"lt_to_not_le/a ? ? Hlt)) applyS a href="cic:/matita/arithmetics/nat/monotonic_le_times_r.def(8)"monotonic_le_times_r/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/not_lt_to_le.def(6)"not_lt_to_le/a/span/span/ qed. -theorem lt_times_n_to_lt_r: +img class="anchor" src="icons/tick.png" id="lt_times_n_to_lt_r"theorem lt_times_n_to_lt_r: ∀n,p,q:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. na title="natural times" href="cic:/fakeuri.def(1)"*/ap a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a na title="natural times" href="cic:/fakeuri.def(1)"*/aq → p a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a q. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/lt_times_n_to_lt_l.def(9)"lt_times_n_to_lt_l/a/span/span/ qed. @@ -767,7 +767,7 @@ qed. *) (************************** minus ******************************) -let rec minus n m ≝ +img class="anchor" src="icons/tick.png" id="minus"let rec minus n m ≝ match n with [ O ⇒ a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a | S p ⇒ @@ -777,22 +777,22 @@ let rec minus n m ≝ interpretation "natural minus" 'minus x y = (minus x y). -theorem minus_S_S: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural minus" href="cic:/fakeuri.def(1)"-/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n a title="natural minus" href="cic:/fakeuri.def(1)"-/am. +img class="anchor" src="icons/tick.png" id="minus_S_S"theorem minus_S_S: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural minus" href="cic:/fakeuri.def(1)"-/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n a title="natural minus" href="cic:/fakeuri.def(1)"-/am. // qed. -theorem minus_O_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="natural minus" href="cic:/fakeuri.def(1)"-/an. +img class="anchor" src="icons/tick.png" id="minus_O_n"theorem minus_O_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="natural minus" href="cic:/fakeuri.def(1)"-/an. #n (cases n) // qed. -theorem minus_n_O: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.na title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ana title="natural minus" href="cic:/fakeuri.def(1)"-/aa href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a. +img class="anchor" src="icons/tick.png" id="minus_n_O"theorem minus_n_O: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.na title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ana title="natural minus" href="cic:/fakeuri.def(1)"-/aa href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a. #n (cases n) // qed. -theorem minus_n_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ana title="natural minus" href="cic:/fakeuri.def(1)"-/an. +img class="anchor" src="icons/tick.png" id="minus_n_n"theorem minus_n_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ana title="natural minus" href="cic:/fakeuri.def(1)"-/an. #n (elim n) // qed. -theorem minus_Sn_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n)a title="natural minus" href="cic:/fakeuri.def(1)"-/an. +img class="anchor" src="icons/tick.png" id="minus_Sn_n"theorem minus_Sn_n: ∀n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n)a title="natural minus" href="cic:/fakeuri.def(1)"-/an. #n (elim n) normalize // qed. -theorem minus_Sn_m: ∀m,n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a (na title="natural minus" href="cic:/fakeuri.def(1)"-/am). +img class="anchor" src="icons/tick.png" id="minus_Sn_m"theorem minus_Sn_m: ∀m,n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a n a title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a (na title="natural minus" href="cic:/fakeuri.def(1)"-/am). (* qualcosa da capire qui #n #m #lenm nelim lenm napplyS refl_eq. *) @a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"nat_elim2/a @@ -810,11 +810,11 @@ theorem not_eq_to_le_to_le_minus: napplyS (not_eq_to_le_to_lt n (S m) H H1) qed. *) -theorem eq_minus_S_pred: ∀n,m. n a title="natural minus" href="cic:/fakeuri.def(1)"-/a (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a(n a title="natural minus" href="cic:/fakeuri.def(1)"-/am). +img class="anchor" src="icons/tick.png" id="eq_minus_S_pred"theorem eq_minus_S_pred: ∀n,m. n a title="natural minus" href="cic:/fakeuri.def(1)"-/a (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a(n a title="natural minus" href="cic:/fakeuri.def(1)"-/am). @a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"nat_elim2/a normalize // qed. -theorem plus_minus: +img class="anchor" src="icons/tick.png" id="plus_minus"theorem plus_minus: ∀m,n,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → (na title="natural minus" href="cic:/fakeuri.def(1)"-/am)a title="natural plus" href="cic:/fakeuri.def(1)"+/ap a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (na title="natural plus" href="cic:/fakeuri.def(1)"+/ap)a title="natural minus" href="cic:/fakeuri.def(1)"-/am. @a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"nat_elim2/a [// @@ -823,27 +823,27 @@ theorem plus_minus: ] qed. -theorem minus_plus_m_m: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (na title="natural plus" href="cic:/fakeuri.def(1)"+/am)a title="natural minus" href="cic:/fakeuri.def(1)"-/am. +img class="anchor" src="icons/tick.png" id="minus_plus_m_m"theorem minus_plus_m_m: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (na title="natural plus" href="cic:/fakeuri.def(1)"+/am)a title="natural minus" href="cic:/fakeuri.def(1)"-/am. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le.con(0,1,1)"le_n/a, a href="cic:/matita/arithmetics/nat/plus_minus.def(5)"plus_minus/a/span/span/ qed. -theorem plus_minus_m_m: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. +img class="anchor" src="icons/tick.png" id="plus_minus_m_m"theorem plus_minus_m_m: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (na title="natural minus" href="cic:/fakeuri.def(1)"-/am)a title="natural plus" href="cic:/fakeuri.def(1)"+/am. #n #m #lemn @a href="cic:/matita/basics/logic/sym_eq.def(2)"sym_eq/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/plus_minus.def(5)"plus_minus/a/span/span/ qed. -theorem le_plus_minus_m_m: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a (na title="natural minus" href="cic:/fakeuri.def(1)"-/am)a title="natural plus" href="cic:/fakeuri.def(1)"+/am. +img class="anchor" src="icons/tick.png" id="le_plus_minus_m_m"theorem le_plus_minus_m_m: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a (na title="natural minus" href="cic:/fakeuri.def(1)"-/am)a title="natural plus" href="cic:/fakeuri.def(1)"+/am. #n (elim n) // #a #Hind #m (cases m) // normalize #n/span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_S_S.def(2)"le_S_S/a/span/span/ qed. -theorem minus_to_plus :∀n,m,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. +img class="anchor" src="icons/tick.png" id="minus_to_plus"theorem minus_to_plus :∀n,m,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → na title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a p → n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a ma title="natural plus" href="cic:/fakeuri.def(1)"+/ap. #n #m #p #lemn #eqp (applyS a href="cic:/matita/arithmetics/nat/plus_minus_m_m.def(7)"plus_minus_m_m/a) // qed. -theorem plus_to_minus :∀n,m,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a ma title="natural plus" href="cic:/fakeuri.def(1)"+/ap → na title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a p. +img class="anchor" src="icons/tick.png" id="plus_to_minus"theorem plus_to_minus :∀n,m,p:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a.n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a ma title="natural plus" href="cic:/fakeuri.def(1)"+/ap → na title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a p. #n #m #p #eqp @a href="cic:/matita/basics/logic/sym_eq.def(2)"sym_eq/a (applyS (a href="cic:/matita/arithmetics/nat/minus_plus_m_m.def(6)"minus_plus_m_m/a p m)) qed. -theorem minus_pred_pred : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n → a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → +img class="anchor" src="icons/tick.png" id="minus_pred_pred"theorem minus_pred_pred : ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n → a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a m → a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a n a title="natural minus" href="cic:/fakeuri.def(1)"-/a a href="cic:/matita/arithmetics/nat/pred.def(1)"pred/a m a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n a title="natural minus" href="cic:/fakeuri.def(1)"-/a m. #n #m #posn #posm @(a href="cic:/matita/arithmetics/nat/lt_O_n_elim.def(4)"lt_O_n_elim/a n posn) @(a href="cic:/matita/arithmetics/nat/lt_O_n_elim.def(4)"lt_O_n_elim/a m posm) //. qed. @@ -900,7 +900,7 @@ qed. (* monotonicity and galois *) -theorem monotonic_le_minus_l: +img class="anchor" src="icons/tick.png" id="monotonic_le_minus_l"theorem monotonic_le_minus_l: ∀p,q,n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. q a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a p → qa title="natural minus" href="cic:/fakeuri.def(1)"-/an a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a pa title="natural minus" href="cic:/fakeuri.def(1)"-/an. @a href="cic:/matita/arithmetics/nat/nat_elim2.def(2)"nat_elim2/a #p #q [#lePO @(a href="cic:/matita/arithmetics/nat/le_n_O_elim.def(4)"le_n_O_elim/a ? lePO) // @@ -909,67 +909,67 @@ theorem monotonic_le_minus_l: ] qed. -theorem le_minus_to_plus: ∀n,m,p. na title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a p → na title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a pa title="natural plus" href="cic:/fakeuri.def(1)"+/am. +img class="anchor" src="icons/tick.png" id="le_minus_to_plus"theorem le_minus_to_plus: ∀n,m,p. na title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a p → na title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a pa title="natural plus" href="cic:/fakeuri.def(1)"+/am. #n #m #p #lep @a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a [|@a href="cic:/matita/arithmetics/nat/le_plus_minus_m_m.def(6)"le_plus_minus_m_m/a | @a href="cic:/matita/arithmetics/nat/monotonic_le_plus_l.def(6)"monotonic_le_plus_l/a // ] qed. -theorem le_minus_to_plus_r: ∀a,b,c. c a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a b → a a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a b a title="natural minus" href="cic:/fakeuri.def(1)"-/a c → a a title="natural plus" href="cic:/fakeuri.def(1)"+/a c a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a b. -#a #b #c #Hlecb #H >(a href="cic:/matita/arithmetics/nat/plus_minus_m_m.def(7)"plus_minus_m_m/a … Hlecb) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_minus_to_plus.def(10)"le_minus_to_plus/a/span/span/ +img class="anchor" src="icons/tick.png" id="le_minus_to_plus_r"theorem le_minus_to_plus_r: ∀a,b,c. c a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a b → a a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a b a title="natural minus" href="cic:/fakeuri.def(1)"-/a c → a a title="natural plus" href="cic:/fakeuri.def(1)"+/a c a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a b. +#a #b #c #Hlecb #H >(a href="cic:/matita/arithmetics/nat/plus_minus_m_m.def(7)"plus_minus_m_m/a … Hlecb) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_minus_to_plus.def(7)"le_minus_to_plus/a/span/span/ qed. -theorem le_plus_to_minus: ∀n,m,p. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a pa title="natural plus" href="cic:/fakeuri.def(1)"+/am → na title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a p. -#n #m #p #lep /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_minus_l.def(9)"monotonic_le_minus_l/a/span/span/ qed. +img class="anchor" src="icons/tick.png" id="le_plus_to_minus"theorem le_plus_to_minus: ∀n,m,p. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a pa title="natural plus" href="cic:/fakeuri.def(1)"+/am → na title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a p. +#n #m #p #lep /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_minus_l.def(5)"monotonic_le_minus_l/a/span/span/ qed. -theorem le_plus_to_minus_r: ∀a,b,c. a a title="natural plus" href="cic:/fakeuri.def(1)"+/a b a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a c → a a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a c a title="natural minus" href="cic:/fakeuri.def(1)"-/ab. +img class="anchor" src="icons/tick.png" id="le_plus_to_minus_r"theorem le_plus_to_minus_r: ∀a,b,c. a a title="natural plus" href="cic:/fakeuri.def(1)"+/a b a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a c → a a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a c a title="natural minus" href="cic:/fakeuri.def(1)"-/ab. #a #b #c #H @(a href="cic:/matita/arithmetics/nat/le_plus_to_le_r.def(6)"le_plus_to_le_r/a … b) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a/span/span/ qed. -theorem lt_minus_to_plus: ∀a,b,c. a a title="natural minus" href="cic:/fakeuri.def(1)"-/a b a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a c → a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a c a title="natural plus" href="cic:/fakeuri.def(1)"+/a b. +img class="anchor" src="icons/tick.png" id="lt_minus_to_plus"theorem lt_minus_to_plus: ∀a,b,c. a a title="natural minus" href="cic:/fakeuri.def(1)"-/a b a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a c → a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a c a title="natural plus" href="cic:/fakeuri.def(1)"+/a b. #a #b #c #H @a href="cic:/matita/arithmetics/nat/not_le_to_lt.def(5)"not_le_to_lt/a @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … (a href="cic:/matita/arithmetics/nat/lt_to_not_le.def(7)"lt_to_not_le/a …H)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_plus_to_minus_r.def(7)"le_plus_to_minus_r/a/span/span/ qed. -theorem lt_minus_to_plus_r: ∀a,b,c. a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a b a title="natural minus" href="cic:/fakeuri.def(1)"-/a c → a a title="natural plus" href="cic:/fakeuri.def(1)"+/a c a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a b. -#a #b #c #H @a href="cic:/matita/arithmetics/nat/not_le_to_lt.def(5)"not_le_to_lt/a @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … (a href="cic:/matita/arithmetics/nat/le_plus_to_minus.def(10)"le_plus_to_minus/a …)) +img class="anchor" src="icons/tick.png" id="lt_minus_to_plus_r"theorem lt_minus_to_plus_r: ∀a,b,c. a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a b a title="natural minus" href="cic:/fakeuri.def(1)"-/a c → a a title="natural plus" href="cic:/fakeuri.def(1)"+/a c a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a b. +#a #b #c #H @a href="cic:/matita/arithmetics/nat/not_le_to_lt.def(5)"not_le_to_lt/a @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … (a href="cic:/matita/arithmetics/nat/le_plus_to_minus.def(7)"le_plus_to_minus/a …)) @a href="cic:/matita/arithmetics/nat/lt_to_not_le.def(7)"lt_to_not_le/a // qed. -theorem lt_plus_to_minus: ∀n,m,p. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a pa title="natural plus" href="cic:/fakeuri.def(1)"+/am → na title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a p. -#n #m #p #lenm #H normalize <a href="cic:/matita/arithmetics/nat/minus_Sn_m.def(5)"minus_Sn_m/a // @a href="cic:/matita/arithmetics/nat/le_plus_to_minus.def(10)"le_plus_to_minus/a // +img class="anchor" src="icons/tick.png" id="lt_plus_to_minus"theorem lt_plus_to_minus: ∀n,m,p. m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → n a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a pa title="natural plus" href="cic:/fakeuri.def(1)"+/am → na title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a p. +#n #m #p #lenm #H normalize <a href="cic:/matita/arithmetics/nat/minus_Sn_m.def(5)"minus_Sn_m/a // @a href="cic:/matita/arithmetics/nat/le_plus_to_minus.def(7)"le_plus_to_minus/a // qed. -theorem lt_plus_to_minus_r: ∀a,b,c. a a title="natural plus" href="cic:/fakeuri.def(1)"+/a b a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a c → a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a c a title="natural minus" href="cic:/fakeuri.def(1)"-/a b. +img class="anchor" src="icons/tick.png" id="lt_plus_to_minus_r"theorem lt_plus_to_minus_r: ∀a,b,c. a a title="natural plus" href="cic:/fakeuri.def(1)"+/a b a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a c → a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a c a title="natural minus" href="cic:/fakeuri.def(1)"-/a b. #a #b #c #H @a href="cic:/matita/arithmetics/nat/le_plus_to_minus_r.def(7)"le_plus_to_minus_r/a // qed. -theorem monotonic_le_minus_r: +img class="anchor" src="icons/tick.png" id="monotonic_le_minus_r"theorem monotonic_le_minus_r: ∀p,q,n:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. q a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a p → na title="natural minus" href="cic:/fakeuri.def(1)"-/ap a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a na title="natural minus" href="cic:/fakeuri.def(1)"-/aq. -#p #q #n #lepq @a href="cic:/matita/arithmetics/nat/le_plus_to_minus.def(10)"le_plus_to_minus/a -@(a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a … (a href="cic:/matita/arithmetics/nat/le_plus_minus_m_m.def(9)"le_plus_minus_m_m/a ? q)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_plus_r.def(3)"monotonic_le_plus_r/a/span/span/ +#p #q #n #lepq @a href="cic:/matita/arithmetics/nat/le_plus_to_minus.def(7)"le_plus_to_minus/a +@(a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a … (a href="cic:/matita/arithmetics/nat/le_plus_minus_m_m.def(6)"le_plus_minus_m_m/a ? q)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_plus_r.def(3)"monotonic_le_plus_r/a/span/span/ qed. -theorem eq_minus_O: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. +img class="anchor" src="icons/tick.png" id="eq_minus_O"theorem eq_minus_O: ∀n,m:a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"nat/a. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → na title="natural minus" href="cic:/fakeuri.def(1)"-/am a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a. -#n #m #lenm @(a href="cic:/matita/arithmetics/nat/le_n_O_elim.def(4)"le_n_O_elim/a (na title="natural minus" href="cic:/fakeuri.def(1)"-/am)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_minus_r.def(11)"monotonic_le_minus_r/a/span/span/ +#n #m #lenm @(a href="cic:/matita/arithmetics/nat/le_n_O_elim.def(4)"le_n_O_elim/a (na title="natural minus" href="cic:/fakeuri.def(1)"-/am)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/monotonic_le_minus_r.def(8)"monotonic_le_minus_r/a/span/span/ qed. -theorem distributive_times_minus: a href="cic:/matita/basics/relations/distributive.def(1)"distributive/a ? a href="cic:/matita/arithmetics/nat/times.fix(0,0,2)"times/a a href="cic:/matita/arithmetics/nat/minus.fix(0,0,1)"minus/a. +img class="anchor" src="icons/tick.png" id="distributive_times_minus"theorem distributive_times_minus: a href="cic:/matita/basics/relations/distributive.def(1)"distributive/a ? a href="cic:/matita/arithmetics/nat/times.fix(0,0,2)"times/a a href="cic:/matita/arithmetics/nat/minus.fix(0,0,1)"minus/a. #a #b #c (cases (a href="cic:/matita/arithmetics/nat/decidable_lt.def(7)"decidable_lt/a b c)) #Hbc - [> a href="cic:/matita/arithmetics/nat/eq_minus_O.def(12)"eq_minus_O/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_plus_b.def(8)"le_plus_b/a/span/span/ >a href="cic:/matita/arithmetics/nat/eq_minus_O.def(12)"eq_minus_O/a // + [> a href="cic:/matita/arithmetics/nat/eq_minus_O.def(9)"eq_minus_O/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_plus_b.def(8)"le_plus_b/a/span/span/ >a href="cic:/matita/arithmetics/nat/eq_minus_O.def(9)"eq_minus_O/a // @a href="cic:/matita/arithmetics/nat/monotonic_le_times_r.def(8)"monotonic_le_times_r/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_plus_b.def(8)"le_plus_b/a/span/span/ |@a href="cic:/matita/basics/logic/sym_eq.def(2)"sym_eq/a (applyS a href="cic:/matita/arithmetics/nat/plus_to_minus.def(7)"plus_to_minus/a) <a href="cic:/matita/arithmetics/nat/distributive_times_plus.def(7)"distributive_times_plus/a @a href="cic:/matita/basics/logic/eq_f.def(3)"eq_f/a (applyS a href="cic:/matita/arithmetics/nat/plus_minus_m_m.def(7)"plus_minus_m_m/a) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/not_lt_to_le.def(6)"not_lt_to_le/a/span/span/ qed. -theorem minus_plus: ∀n,m,p. na title="natural minus" href="cic:/fakeuri.def(1)"-/ama title="natural minus" href="cic:/fakeuri.def(1)"-/ap a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n a title="natural minus" href="cic:/fakeuri.def(1)"-/a(ma title="natural plus" href="cic:/fakeuri.def(1)"+/ap). +img class="anchor" src="icons/tick.png" id="minus_plus"theorem minus_plus: ∀n,m,p. na title="natural minus" href="cic:/fakeuri.def(1)"-/ama title="natural minus" href="cic:/fakeuri.def(1)"-/ap a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a n a title="natural minus" href="cic:/fakeuri.def(1)"-/a(ma title="natural plus" href="cic:/fakeuri.def(1)"+/ap). #n #m #p cases (a href="cic:/matita/arithmetics/nat/decidable_le.def(6)"decidable_le/a (ma title="natural plus" href="cic:/fakeuri.def(1)"+/ap) n) #Hlt [@a href="cic:/matita/arithmetics/nat/plus_to_minus.def(7)"plus_to_minus/a @a href="cic:/matita/arithmetics/nat/plus_to_minus.def(7)"plus_to_minus/a <a href="cic:/matita/arithmetics/nat/associative_plus.def(4)"associative_plus/a @a href="cic:/matita/arithmetics/nat/minus_to_plus.def(8)"minus_to_plus/a // |cut (n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a ma title="natural plus" href="cic:/fakeuri.def(1)"+/ap) [@(a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a … (a href="cic:/matita/arithmetics/nat/le_n_Sn.def(1)"le_n_Sn/a …)) @a href="cic:/matita/arithmetics/nat/not_le_to_lt.def(5)"not_le_to_lt/a //] - #H >a href="cic:/matita/arithmetics/nat/eq_minus_O.def(12)"eq_minus_O/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/eq_minus_O.def(12)"eq_minus_O/a, a href="cic:/matita/arithmetics/nat/monotonic_le_minus_l.def(9)"monotonic_le_minus_l/a/span/span/ + #H >a href="cic:/matita/arithmetics/nat/eq_minus_O.def(9)"eq_minus_O/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/eq_minus_O.def(9)"eq_minus_O/a, a href="cic:/matita/arithmetics/nat/monotonic_le_minus_l.def(5)"monotonic_le_minus_l/a/span/span/ ] qed. @@ -979,7 +979,7 @@ theorem plus_minus: ∀n,m,p. p ≤ m → (n+m)-p = n +(m-p). >associative_plus (a href="cic:/matita/arithmetics/nat/le_to_leb_true.def(7)"le_to_leb_true/a …) // @(a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a ? (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/not_le_to_lt.def(5)"not_le_to_lt/a/span/span/ ] qed. -lemma le_minr: ∀i,n,m. i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/min.def(2)"min/a n m → i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. +img class="anchor" src="icons/tick.png" id="le_minr"lemma le_minr: ∀i,n,m. i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/min.def(2)"min/a n m → i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m. #i #n #m normalize @a href="cic:/matita/arithmetics/nat/leb_elim.def(6)"leb_elim/a normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a/span/span/ qed. -lemma le_minl: ∀i,n,m. i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/min.def(2)"min/a n m → i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. +img class="anchor" src="icons/tick.png" id="le_minl"lemma le_minl: ∀i,n,m. i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/min.def(2)"min/a n m → i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_minr.def(7)"le_minr/a/span/span/ qed. -lemma to_min: ∀i,n,m. i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/min.def(2)"min/a n m. +img class="anchor" src="icons/tick.png" id="to_min"lemma to_min: ∀i,n,m. i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a n → i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a m → i a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a href="cic:/matita/arithmetics/nat/min.def(2)"min/a n m. #i #n #m #lein #leim normalize (cases (a href="cic:/matita/arithmetics/nat/leb.fix(0,0,1)"leb/a n m)) normalize // qed. -lemma commutative_max: a href="cic:/matita/basics/relations/commutative.def(1)"commutative/a ? a href="cic:/matita/arithmetics/nat/max.def(2)"max/a. +img class="anchor" src="icons/tick.png" id="commutative_max"lemma commutative_max: a href="cic:/matita/basics/relations/commutative.def(1)"commutative/a ? a href="cic:/matita/arithmetics/nat/max.def(2)"max/a. #n #m normalize @a href="cic:/matita/arithmetics/nat/leb_elim.def(6)"leb_elim/a [@a href="cic:/matita/arithmetics/nat/leb_elim.def(6)"leb_elim/a normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_to_le_to_eq.def(5)"le_to_le_to_eq/a/span/span/ |#notle >(a href="cic:/matita/arithmetics/nat/le_to_leb_true.def(7)"le_to_leb_true/a …) // @(a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a ? (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m)) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/not_le_to_lt.def(5)"not_le_to_lt/a/span/span/ ] qed. -lemma le_maxl: ∀i,n,m. a href="cic:/matita/arithmetics/nat/max.def(2)"max/a n m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i. +img class="anchor" src="icons/tick.png" id="le_maxl"lemma le_maxl: ∀i,n,m. a href="cic:/matita/arithmetics/nat/max.def(2)"max/a n m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i → n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i. #i #n #m normalize @a href="cic:/matita/arithmetics/nat/leb_elim.def(6)"leb_elim/a normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/transitive_le.def(3)"transitive_le/a/span/span/ qed. -lemma le_maxr: ∀i,n,m. a href="cic:/matita/arithmetics/nat/max.def(2)"max/a n m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i. +img class="anchor" src="icons/tick.png" id="le_maxr"lemma le_maxr: ∀i,n,m. a href="cic:/matita/arithmetics/nat/max.def(2)"max/a n m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i. /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/arithmetics/nat/le_maxl.def(7)"le_maxl/a/span/span/ qed. -lemma to_max: ∀i,n,m. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i → a href="cic:/matita/arithmetics/nat/max.def(2)"max/a n m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i. +img class="anchor" src="icons/tick.png" id="to_max"lemma to_max: ∀i,n,m. n a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i → m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i → a href="cic:/matita/arithmetics/nat/max.def(2)"max/a n m a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a i. #i #n #m #leni #lemi normalize (cases (a href="cic:/matita/arithmetics/nat/leb.fix(0,0,1)"leb/a n m)) normalize // qed. diff --git a/weblib/basics/bool.ma b/weblib/basics/bool.ma index bd832f5f9..d532c3186 100644 --- a/weblib/basics/bool.ma +++ b/weblib/basics/bool.ma @@ -12,80 +12,80 @@ include "basics/relations.ma". (********** bool **********) -inductive bool: Type[0] ≝ +img class="anchor" src="icons/tick.png" id="bool"inductive bool: Type[0] ≝ | true : bool | false : bool. (* destruct does not work *) -theorem not_eq_true_false : a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. +img class="anchor" src="icons/tick.png" id="not_eq_true_false"theorem not_eq_true_false : a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. @a href="cic:/matita/basics/logic/Not.con(0,1,1)"nmk/a #Heq change with match a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a with [true ⇒ a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a|false ⇒ a href="cic:/matita/basics/logic/True.ind(1,0,0)"True/a] >Heq // qed. -definition notb : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ +img class="anchor" src="icons/tick.png" id="notb"definition notb : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ λ b:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. match b with [true ⇒ a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a|false ⇒ a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a ]. interpretation "boolean not" 'not x = (notb x). -theorem notb_elim: ∀ b:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a.∀ P:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → Prop. +img class="anchor" src="icons/tick.png" id="notb_elim"theorem notb_elim: ∀ b:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a.∀ P:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → Prop. match b with [ true ⇒ P a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a | false ⇒ P a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a] → P (a href="cic:/matita/basics/bool/notb.def(1)"notb/a b). #b #P elim b normalize // qed. -theorem notb_notb: ∀b:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. a href="cic:/matita/basics/bool/notb.def(1)"notb/a (a href="cic:/matita/basics/bool/notb.def(1)"notb/a b) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b. +img class="anchor" src="icons/tick.png" id="notb_notb"theorem notb_notb: ∀b:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. a href="cic:/matita/basics/bool/notb.def(1)"notb/a (a href="cic:/matita/basics/bool/notb.def(1)"notb/a b) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b. #b elim b // qed. -theorem injective_notb: a href="cic:/matita/basics/relations/injective.def(1)"injective/a a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a a href="cic:/matita/basics/bool/notb.def(1)"notb/a. +img class="anchor" src="icons/tick.png" id="injective_notb"theorem injective_notb: a href="cic:/matita/basics/relations/injective.def(1)"injective/a a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a a href="cic:/matita/basics/bool/notb.def(1)"notb/a. #b1 #b2 #H // qed. -theorem noteq_to_eqnot: ∀b1,b2. b1 a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a b2 → b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/notb.def(1)"notb/a b2. +img class="anchor" src="icons/tick.png" id="noteq_to_eqnot"theorem noteq_to_eqnot: ∀b1,b2. b1 a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a b2 → b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/notb.def(1)"notb/a b2. * * // #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ qed. -theorem eqnot_to_noteq: ∀b1,b2. b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/notb.def(1)"notb/a b2 → b1 a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a b2. +img class="anchor" src="icons/tick.png" id="eqnot_to_noteq"theorem eqnot_to_noteq: ∀b1,b2. b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/notb.def(1)"notb/a b2 → b1 a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a b2. * * normalize // #H @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … a href="cic:/matita/basics/bool/not_eq_true_false.def(3)"not_eq_true_false/a) // qed. -definition andb : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ +img class="anchor" src="icons/tick.png" id="andb"definition andb : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ λb1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. match b1 with [ true ⇒ b2 | false ⇒ a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a ]. interpretation "boolean and" 'and x y = (andb x y). -theorem andb_elim: ∀ b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. ∀ P:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → Prop. +img class="anchor" src="icons/tick.png" id="andb_elim"theorem andb_elim: ∀ b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. ∀ P:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → Prop. match b1 with [ true ⇒ P b2 | false ⇒ P a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a] → P (b1 a title="boolean and" href="cic:/fakeuri.def(1)"∧/a b2). #b1 #b2 #P (elim b1) normalize // qed. -theorem andb_true_l: ∀ b1,b2. (b1 a title="boolean and" href="cic:/fakeuri.def(1)"∧/a b2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +img class="anchor" src="icons/tick.png" id="andb_true_l"theorem andb_true_l: ∀ b1,b2. (b1 a title="boolean and" href="cic:/fakeuri.def(1)"∧/a b2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #b1 (cases b1) normalize // qed. -theorem andb_true_r: ∀b1,b2. (b1 a title="boolean and" href="cic:/fakeuri.def(1)"∧/a b2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → b2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +img class="anchor" src="icons/tick.png" id="andb_true_r"theorem andb_true_r: ∀b1,b2. (b1 a title="boolean and" href="cic:/fakeuri.def(1)"∧/a b2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → b2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #b1 #b2 (cases b1) normalize // (cases b2) // qed. -theorem andb_true: ∀b1,b2. (b1 a title="boolean and" href="cic:/fakeuri.def(1)"∧/a b2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="logical and" href="cic:/fakeuri.def(1)"∧/a b2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +img class="anchor" src="icons/tick.png" id="andb_true"theorem andb_true: ∀b1,b2. (b1 a title="boolean and" href="cic:/fakeuri.def(1)"∧/a b2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="logical and" href="cic:/fakeuri.def(1)"∧/a b2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a, a href="cic:/matita/basics/bool/andb_true_r.def(4)"andb_true_r/a, a href="cic:/matita/basics/bool/andb_true_l.def(4)"andb_true_l/a/span/span/ qed. -definition orb : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ +img class="anchor" src="icons/tick.png" id="orb"definition orb : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ λb1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a.match b1 with [ true ⇒ a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a | false ⇒ b2]. interpretation "boolean or" 'or x y = (orb x y). -theorem orb_elim: ∀ b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. ∀ P:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → Prop. +img class="anchor" src="icons/tick.png" id="orb_elim"theorem orb_elim: ∀ b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. ∀ P:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → Prop. match b1 with [ true ⇒ P a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a | false ⇒ P b2] → P (a href="cic:/matita/basics/bool/orb.def(1)"orb/a b1 b2). #b1 #b2 #P (elim b1) normalize // qed. -lemma orb_true_r1: ∀b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. +img class="anchor" src="icons/tick.png" id="orb_true_r1"lemma orb_true_r1: ∀b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → (b1 a title="boolean or" href="cic:/fakeuri.def(1)"∨/a b2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #b1 #b2 #H >H // qed. -lemma orb_true_r2: ∀b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. +img class="anchor" src="icons/tick.png" id="orb_true_r2"lemma orb_true_r2: ∀b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. b2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → (b1 a title="boolean or" href="cic:/fakeuri.def(1)"∨/a b2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #b1 #b2 #H >H cases b1 // qed. -lemma orb_true_l: ∀b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. +img class="anchor" src="icons/tick.png" id="orb_true_l"lemma orb_true_l: ∀b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. (b1 a title="boolean or" href="cic:/fakeuri.def(1)"∨/a b2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → (b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) a title="logical or" href="cic:/fakeuri.def(1)"∨/a (b2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a). * normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ qed. -definition xorb : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ +img class="anchor" src="icons/tick.png" id="xorb"definition xorb : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ λb1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. match b1 with [ true ⇒ match b2 with [ true ⇒ a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a | false ⇒ a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a ] @@ -97,6 +97,11 @@ notation > "'if' term 46 e 'then' term 46 t 'else' term 46 f" non associative wi notation < "hvbox('if' \nbsp term 46 e \nbsp break 'then' \nbsp term 46 t \nbsp break 'else' \nbsp term 49 f \nbsp)" non associative with precedence 46 for @{ match $e with [ true ⇒ $t | false ⇒ $f] }. -definition ite ≝ λA:Type[0].λe.λt,f:A.match e with [ true ⇒ t | false ⇒ f ]. +img class="anchor" src="icons/tick.png" id="bool_to_decidable_eq"theorem bool_to_decidable_eq: + ∀b1,b2:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. a href="cic:/matita/basics/logic/decidable.def(1)"decidable/a (b1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ab2). +#b1 #b2 (cases b1) (cases b2) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ %2 /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/bool/eqnot_to_noteq.def(4)"eqnot_to_noteq/a/span/span/ qed. + +img class="anchor" src="icons/tick.png" id="true_or_false"theorem true_or_false: +∀b:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. b a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="logical or" href="cic:/fakeuri.def(1)"∨/a b a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. +#b (cases b) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ qed. -interpretation "if then else" 'ite e t f = (ite ? e t f). diff --git a/weblib/basics/jmeq.ma b/weblib/basics/jmeq.ma index ad3aa9371..7838f0e97 100644 --- a/weblib/basics/jmeq.ma +++ b/weblib/basics/jmeq.ma @@ -11,82 +11,82 @@ include "basics/logic.ma". -inductive Sigma: Type[1] ≝ +img class="anchor" src="icons/tick.png" id="Sigma"inductive Sigma: Type[1] ≝ mk_Sigma: ∀p1: Type[0]. p1 → Sigma. -definition p1: Sigma → Type[0]. +img class="anchor" src="icons/tick.png" id="p1"definition p1: a href="cic:/matita/basics/jmeq/Sigma.ind(1,0,0)"Sigma/a → Type[0]. #S cases S #Y #_ @Y qed. -definition p2: ∀S:Sigma. p1 S. +img class="anchor" src="icons/tick.png" id="p2"definition p2: ∀S:a href="cic:/matita/basics/jmeq/Sigma.ind(1,0,0)"Sigma/a. a href="cic:/matita/basics/jmeq/p1.def(1)"p1/a S. #S cases S #Y #x @x qed. -inductive jmeq (A:Type[0]) (x:A) : ∀B:Type[0]. B →Prop ≝ +img class="anchor" src="icons/tick.png" id="jmeq"inductive jmeq (A:Type[0]) (x:A) : ∀B:Type[0]. B →Prop ≝ jmrefl : jmeq A x A x. -definition eqProp ≝ λA:Prop.eq A. +img class="anchor" src="icons/tick.png" id="eqProp"definition eqProp ≝ λA:Prop.a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a A. -lemma K : ∀A.∀x:A.∀h:x=x. eqProp ? h (refl A x). -#A #x #h @(refl ? h: eqProp ? ? ?). +img class="anchor" src="icons/tick.png" id="K"lemma K : ∀A.∀x:A.∀h:xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax. a href="cic:/matita/basics/jmeq/eqProp.def(1)" title="null"eqProp/a ? h (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A x). +#A #x #h @(a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? h: a href="cic:/matita/basics/jmeq/eqProp.def(1)"eqProp/a ? ? ?). qed. -definition cast: - ∀A,B:Type[0].∀E:A=B. A → B. +img class="anchor" src="icons/tick.png" id="cast"definition cast: + ∀A,B:Type[0].∀E:Aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aB. A → B. #A #B #E cases E #X @X qed. -lemma tech1: - ∀Aa,Bb:Sigma.∀E:Aa=Bb. - cast (p1 Aa) (p1 Bb) ? (p2 Aa) = p2 Bb. +img class="anchor" src="icons/tick.png" id="tech1"lemma tech1: + ∀Aa,Bb:a href="cic:/matita/basics/jmeq/Sigma.ind(1,0,0)"Sigma/a.∀E:Aaa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aBb. + a href="cic:/matita/basics/jmeq/cast.def(1)"cast/a (a href="cic:/matita/basics/jmeq/p1.def(1)"p1/a Aa) (a href="cic:/matita/basics/jmeq/p1.def(1)"p1/a Bb) ? (a href="cic:/matita/basics/jmeq/p2.def(2)"p2/a Aa) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/jmeq/p2.def(2)"p2/a Bb. [2: >E % | #Aa #Bb #E >E cases Bb #B #b normalize % ] qed. -lemma gral: ∀A.∀x,y:A. - mk_Sigma A x = mk_Sigma A y → x=y. - #A #x #y #E lapply (tech1 ?? E) - generalize in ⊢ (??(???%?)? → ?) #E1 - normalize in E1; >(K ?? E1) normalize +img class="anchor" src="icons/tick.png" id="gral"lemma gral: ∀A.∀x,y:A. + a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A y → xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ay. + #A #x #y #E lapply (a href="cic:/matita/basics/jmeq/tech1.def(3)"tech1/a ?? E) + generalize in ⊢ (??(???%?)? → ?); #E1 + normalize in E1; >(a href="cic:/matita/basics/jmeq/K.def(2)"K/a ?? E1) normalize #H @H qed. -axiom daemon: False. +img class="anchor" src="icons/tick.png" id="daemon"axiom daemon: a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a. -lemma jm_to_eq_sigma: - ∀A,x,y. jmeq A x A y → mk_Sigma A x = mk_Sigma A y. +img class="anchor" src="icons/tick.png" id="jm_to_eq_sigma"lemma jm_to_eq_sigma: + ∀A,x,y. a href="cic:/matita/basics/jmeq/jmeq.ind(1,0,2)"jmeq/a A x A y → a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A y. #A #x #y #E cases E in ⊢ (???%); % qed. -definition curry: +img class="anchor" src="icons/tick.png" id="curry"definition curry: ∀A,x. - (∀y. jmeq A x A y → Type[0]) → - (∀y. mk_Sigma A x = mk_Sigma A y → Type[0]). - #A #x #f #y #E @(f y) >(gral ??? E) % + (∀y. a href="cic:/matita/basics/jmeq/jmeq.ind(1,0,2)"jmeq/a A x A y → Type[0]) → + (∀y. a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A y → Type[0]). + #A #x #f #y #E @(f y) >(a href="cic:/matita/basics/jmeq/gral.def(4)"gral/a ??? E) % qed. -lemma G : ∀A.∀x:A.∀P:∀y:A.mk_Sigma A x = mk_Sigma A y →Type[0]. - P x (refl ? (mk_Sigma A x)) → ∀y.∀h:mk_Sigma A x = mk_Sigma A y. +img class="anchor" src="icons/tick.png" id="G"lemma G : ∀A.∀x:A.∀P:∀y:A.a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A y →Type[0]. + P x (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? (a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A x)) → ∀y.∀h:a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A y. P y h. - #A #x #P #H #y #E lapply (gral ??? E) #G generalize in match E; + #A #x #P #H #y #E lapply (a href="cic:/matita/basics/jmeq/gral.def(4)"gral/a ??? E) #G generalize in match E; @(match G - return λy.λ_. ∀e:mk_Sigma A x = mk_Sigma A y. P y e + return λy.λ_. ∀e:a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/jmeq/Sigma.con(0,1,0)"mk_Sigma/a A y. P y e with [refl ⇒ ?]) - #E <(sym_eq ??? (K ?? E)) @H + #E <(a href="cic:/matita/basics/logic/sym_eq.def(2)"sym_eq/a ??? (a href="cic:/matita/basics/jmeq/K.def(2)"K/a ?? E)) @H qed. -definition PP: ∀A:Prop.∀P:A → Type[0]. A → Type[0]. +img class="anchor" src="icons/tick.png" id="PP"definition PP: ∀A:Prop.∀P:A → Type[0]. A → Type[0]. #A #P #a @(P a) qed. -lemma E : ∀A.∀x:A.∀P:∀y:A.jmeq A x A y→Type[0]. - PP ? (P x) (jmrefl A x) → ∀y.∀h:jmeq A x A y.PP ? (P y) h. - #A #a #P #H #b #E letin F ≝ (jm_to_eq_sigma ??? E) - lapply (G ?? (curry ?? P) ?? F) +img class="anchor" src="icons/tick.png" id="E"lemma E : ∀A.∀x:A.∀P:∀y:A.a href="cic:/matita/basics/jmeq/jmeq.ind(1,0,2)"jmeq/a A x A y→Type[0]. + a href="cic:/matita/basics/jmeq/PP.def(1)"PP/a ? (P x) (a href="cic:/matita/basics/jmeq/jmeq.con(0,1,2)"jmrefl/a A x) → ∀y.∀h:a href="cic:/matita/basics/jmeq/jmeq.ind(1,0,2)"jmeq/a A x A y.a href="cic:/matita/basics/jmeq/PP.def(1)"PP/a ? (P y) h. + #A #a #P #H #b #E letin F ≝ (a href="cic:/matita/basics/jmeq/jm_to_eq_sigma.def(1)"jm_to_eq_sigma/a ??? E) + lapply (a href="cic:/matita/basics/jmeq/G.def(5)"G/a ?? (a href="cic:/matita/basics/jmeq/curry.def(5)"curry/a ?? P) ?? F) [ normalize // - | #H whd in H; @(H : PP ? (P b) ?) ] + | #H whd in H; @(H : a href="cic:/matita/basics/jmeq/PP.def(1)"PP/a ? (P b) ?) ] qed. -lemma jmeq_elim : ∀A.∀x:A.∀P:∀y:A.jmeq A x A y→Type[0]. - P x (jmrefl A x) → ∀y.∀h:jmeq A x A y.P y h ≝ E. +img class="anchor" src="icons/tick.png" id="jmeq_elim"lemma jmeq_elim : ∀A.∀x:A.∀P:∀y:A.a href="cic:/matita/basics/jmeq/jmeq.ind(1,0,2)"jmeq/a A x A y→Type[0]. + P x (a href="cic:/matita/basics/jmeq/jmeq.con(0,1,2)"jmrefl/a A x) → ∀y.∀h:a href="cic:/matita/basics/jmeq/jmeq.ind(1,0,2)"jmeq/a A x A y.P y h ≝ a href="cic:/matita/basics/jmeq/E.def(6)"E/a. \ No newline at end of file diff --git a/weblib/basics/list.ma b/weblib/basics/list.ma index 6dcc181f5..d542416c4 100644 --- a/weblib/basics/list.ma +++ b/weblib/basics/list.ma @@ -11,7 +11,7 @@ include "arithmetics/nat.ma". -inductive list (A:Type[0]) : Type[0] := +img class="anchor" src="icons/tick.png" id="list"inductive list (A:Type[0]) : Type[0] := | nil: list A | cons: A -> list A -> list A. @@ -30,12 +30,12 @@ notation "hvbox(l1 break @ l2)" interpretation "nil" 'nil = (nil ?). interpretation "cons" 'cons hd tl = (cons ? hd tl). -definition not_nil: ∀A:Type[0].a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A → Prop ≝ +img class="anchor" src="icons/tick.png" id="not_nil"definition not_nil: ∀A:Type[0].a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A → Prop ≝ λA.λl.match l with [ nil ⇒ a href="cic:/matita/basics/logic/True.ind(1,0,0)"True/a | cons hd tl ⇒ a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a ]. -theorem nil_cons: - ∀A:Type[0].∀l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.∀a:A. aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a [a title="nil" href="cic:/fakeuri.def(1)"]/a. - #A #l #a @a href="cic:/matita/basics/logic/Not.con(0,1,1)"nmk/a #Heq (change with (a href="cic:/matita/basics/list/not_nil.def(1)"not_nil/a ? (a:a title="cons" href="cic:/fakeuri.def(1)":/al))) >Heq // +img class="anchor" src="icons/tick.png" id="nil_cons"theorem nil_cons: + ∀A:Type[0].∀l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.∀a:A. aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a. + #A #l #a @a href="cic:/matita/basics/logic/Not.con(0,1,1)"nmk/a #Heq (change with (a href="cic:/matita/basics/list/not_nil.def(1)"not_nil/a ? (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al))) >Heq // qed. (* @@ -44,23 +44,23 @@ let rec id_list A (l: list A) on l := [ nil => [] | (cons hd tl) => hd :: id_list A tl ]. *) -let rec append A (l1: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) l2 on l1 ≝ +img class="anchor" src="icons/tick.png" id="append"let rec append A (l1: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) l2 on l1 ≝ match l1 with [ nil ⇒ l2 - | cons hd tl ⇒ hd :a title="cons" href="cic:/fakeuri.def(1)":/a append A tl l2 ]. + | cons hd tl ⇒ hd a title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/a append A tl l2 ]. -definition hd ≝ λA.λl: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.λd:A. +img class="anchor" src="icons/tick.png" id="hd"definition hd ≝ λA.λl: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.λd:A. match l with [ nil ⇒ d | cons a _ ⇒ a]. -definition tail ≝ λA.λl: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A. - match l with [ nil ⇒ [a title="nil" href="cic:/fakeuri.def(1)"]/a | cons hd tl ⇒ tl]. +img class="anchor" src="icons/tick.png" id="tail"definition tail ≝ λA.λl: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A. + match l with [ nil ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a | cons hd tl ⇒ tl]. interpretation "append" 'append l1 l2 = (append ? l1 l2). -theorem append_nil: ∀A.∀l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.l a title="append" href="cic:/fakeuri.def(1)"@/a [a title="nil" href="cic:/fakeuri.def(1)"]/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a l. +img class="anchor" src="icons/tick.png" id="append_nil"theorem append_nil: ∀A.∀l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.l a title="append" href="cic:/fakeuri.def(1)"@/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a l. #A #l (elim l) normalize // qed. -theorem associative_append: +img class="anchor" src="icons/tick.png" id="associative_append"theorem associative_append: ∀A.a href="cic:/matita/basics/relations/associative.def(1)"associative/a (a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) (a href="cic:/matita/basics/list/append.fix(0,1,1)"append/a A). #A #l1 #l2 #l3 (elim l1) normalize // qed. @@ -70,51 +70,51 @@ ntheorem cons_append_commute: a :: (l1 @ l2) = (a :: l1) @ l2. //; nqed. *) -theorem append_cons:∀A.∀a:A.∀l,l1.la title="append" href="cic:/fakeuri.def(1)"@/a(a:a title="cons" href="cic:/fakeuri.def(1)":/al1)a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a(la title="append" href="cic:/fakeuri.def(1)"@/a(a:a title="cons" href="cic:/fakeuri.def(1)":/a[a title="nil" href="cic:/fakeuri.def(1)"]/a))a title="append" href="cic:/fakeuri.def(1)"@/al1.span style="text-decoration: underline;"/spanspan class="autotactic"/span +img class="anchor" src="icons/tick.png" id="append_cons"theorem append_cons:∀A.∀a:A.∀l,l1.la title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al1)a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a(la title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a))a title="append" href="cic:/fakeuri.def(1)"@/al1.span style="text-decoration: underline;"/spanspan class="autotactic"/span #A #a #l1 #l2 >a href="cic:/matita/basics/list/associative_append.def(4)"associative_append/a // qed. -theorem nil_append_elim: ∀A.∀l1,l2: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.∀P:?→?→Prop. - l1a title="append" href="cic:/fakeuri.def(1)"@/al2a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a[a title="nil" href="cic:/fakeuri.def(1)"]/a → P (a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a A) (a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a A) → P l1 l2. +img class="anchor" src="icons/tick.png" id="nil_append_elim"theorem nil_append_elim: ∀A.∀l1,l2: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.∀P:?→?→Prop. + l1a title="append" href="cic:/fakeuri.def(1)"@/al2a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a → P (a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a A) (a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a A) → P l1 l2. #A #l1 #l2 #P (cases l1) normalize // #a #l3 #heq destruct qed. -theorem nil_to_nil: ∀A.∀l1,l2:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A. - l1a title="append" href="cic:/fakeuri.def(1)"@/al2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a [a title="nil" href="cic:/fakeuri.def(1)"]/a → l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a [a title="nil" href="cic:/fakeuri.def(1)"]/a a title="logical and" href="cic:/fakeuri.def(1)"∧/a l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a [a title="nil" href="cic:/fakeuri.def(1)"]/a. +img class="anchor" src="icons/tick.png" id="nil_to_nil"theorem nil_to_nil: ∀A.∀l1,l2:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A. + l1a title="append" href="cic:/fakeuri.def(1)"@/al2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a → l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a a title="logical and" href="cic:/fakeuri.def(1)"∧/a l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a. #A #l1 #l2 #isnil @(a href="cic:/matita/basics/list/nil_append_elim.def(4)"nil_append_elim/a A l1 l2) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ qed. (* iterators *) -let rec map (A,B:Type[0]) (f: A → B) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a B ≝ - match l with [ nil ⇒ a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a ? | cons x tl ⇒ f x :a title="cons" href="cic:/fakeuri.def(1)":/a (map A B f tl)]. +img class="anchor" src="icons/tick.png" id="map"let rec map (A,B:Type[0]) (f: A → B) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a B ≝ + match l with [ nil ⇒ a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a ? | cons x tl ⇒ f x a title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/a (map A B f tl)]. -lemma map_append : ∀A,B,f,l1,l2. +img class="anchor" src="icons/tick.png" id="map_append"lemma map_append : ∀A,B,f,l1,l2. (a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a A B f l1) a title="append" href="cic:/fakeuri.def(1)"@/a (a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a A B f l2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a A B f (l1a title="append" href="cic:/fakeuri.def(1)"@/al2). #A #B #f #l1 elim l1 [ #l2 @a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a | #h #t #IH #l2 normalize // ] qed. -let rec foldr (A,B:Type[0]) (f:A → B → B) (b:B) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l :B ≝ +img class="anchor" src="icons/tick.png" id="foldr"let rec foldr (A,B:Type[0]) (f:A → B → B) (b:B) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l :B ≝ match l with [ nil ⇒ b | cons a l ⇒ f a (foldr A B f b l)]. -definition filter ≝ +img class="anchor" src="icons/tick.png" id="filter"definition filter ≝ λT.λp:T → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. - a href="cic:/matita/basics/list/foldr.fix(0,4,1)"foldr/a T (a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a T) (λx,l0.if (p x) then (x:a title="cons" href="cic:/fakeuri.def(1)":/al0) else l0) (a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a T). + a href="cic:/matita/basics/list/foldr.fix(0,4,1)"foldr/a T (a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a T) (λx,l0.if (p x) then (xa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al0) else l0) (a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a T). -definition compose ≝ λA,B,C.λf:A→B→C.λl1,l2. - a href="cic:/matita/basics/list/foldr.fix(0,4,1)"foldr/a ?? (λi,acc.(a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a ?? (f i) l2)a title="append" href="cic:/fakeuri.def(1)"@/aacc) [ a title="nil" href="cic:/fakeuri.def(1)"]/a l1. +img class="anchor" src="icons/tick.png" id="compose"definition compose ≝ λA,B,C.λf:A→B→C.λl1,l2. + a href="cic:/matita/basics/list/foldr.fix(0,4,1)"foldr/a ?? (λi,acc.(a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a ?? (f i) l2)a title="append" href="cic:/fakeuri.def(1)"@/aacc) a title="nil" href="cic:/fakeuri.def(1)"[/a a title="nil" href="cic:/fakeuri.def(1)"]/a l1. -lemma filter_true : ∀A,l,a,p. p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → - a href="cic:/matita/basics/list/filter.def(2)"filter/a A p (a:a title="cons" href="cic:/fakeuri.def(1)":/al) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a :a title="cons" href="cic:/fakeuri.def(1)":/a a href="cic:/matita/basics/list/filter.def(2)"filter/a A p l. +img class="anchor" src="icons/tick.png" id="filter_true"lemma filter_true : ∀A,l,a,p. p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → + a href="cic:/matita/basics/list/filter.def(2)"filter/a A p (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a a title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/a a href="cic:/matita/basics/list/filter.def(2)"filter/a A p l. #A #l #a #p #pa (elim l) normalize >pa normalize // qed. -lemma filter_false : ∀A,l,a,p. p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → - a href="cic:/matita/basics/list/filter.def(2)"filter/a A p (a:a title="cons" href="cic:/fakeuri.def(1)":/al) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/list/filter.def(2)"filter/a A p l. +img class="anchor" src="icons/tick.png" id="filter_false"lemma filter_false : ∀A,l,a,p. p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → + a href="cic:/matita/basics/list/filter.def(2)"filter/a A p (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/list/filter.def(2)"filter/a A p l. #A #l #a #p #pa (elim l) normalize >pa normalize // qed. -theorem eq_map : ∀A,B,f,g,l. (∀x.f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a g x) → a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a A B f l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a A B g l. +img class="anchor" src="icons/tick.png" id="eq_map"theorem eq_map : ∀A,B,f,g,l. (∀x.f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a g x) → a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a A B f l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a A B g l. #A #B #f #g #l #eqfg (elim l) normalize // qed. (* @@ -125,40 +125,40 @@ match l1 with ]. *) (**************************** reverse *****************************) -let rec rev_append S (l1,l2:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S) on l1 ≝ +img class="anchor" src="icons/tick.png" id="rev_append"let rec rev_append S (l1,l2:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S) on l1 ≝ match l1 with [ nil ⇒ l2 - | cons a tl ⇒ rev_append S tl (a:a title="cons" href="cic:/fakeuri.def(1)":/al2) + | cons a tl ⇒ rev_append S tl (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al2) ] . -definition reverse ≝λS.λl.a href="cic:/matita/basics/list/rev_append.fix(0,1,1)"rev_append/a S l [a title="nil" href="cic:/fakeuri.def(1)"]/a. +img class="anchor" src="icons/tick.png" id="reverse"definition reverse ≝λS.λl.a href="cic:/matita/basics/list/rev_append.fix(0,1,1)"rev_append/a S l a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a. -lemma reverse_single : ∀S,a. a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S (a:a title="cons" href="cic:/fakeuri.def(1)":/a[a title="nil" href="cic:/fakeuri.def(1)"]/a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a:a title="cons" href="cic:/fakeuri.def(1)":/a[a title="nil" href="cic:/fakeuri.def(1)"]/a). +img class="anchor" src="icons/tick.png" id="reverse_single"lemma reverse_single : ∀S,a. a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a). // qed. -lemma rev_append_def : ∀S,l1,l2. +img class="anchor" src="icons/tick.png" id="rev_append_def"lemma rev_append_def : ∀S,l1,l2. a href="cic:/matita/basics/list/rev_append.fix(0,1,1)"rev_append/a S l1 l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S l1) a title="append" href="cic:/fakeuri.def(1)"@/a l2 . #S #l1 elim l1 normalize // qed. -lemma reverse_cons : ∀S,a,l. a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S (a:a title="cons" href="cic:/fakeuri.def(1)":/al) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S l)a title="append" href="cic:/fakeuri.def(1)"@/a(a:a title="cons" href="cic:/fakeuri.def(1)":/a[a title="nil" href="cic:/fakeuri.def(1)"]/a). +img class="anchor" src="icons/tick.png" id="reverse_cons"lemma reverse_cons : ∀S,a,l. a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S l)a title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a). #S #a #l whd in ⊢ (??%?); // qed. -lemma reverse_append: ∀S,l1,l2. +img class="anchor" src="icons/tick.png" id="reverse_append"lemma reverse_append: ∀S,l1,l2. a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S (l1 a title="append" href="cic:/fakeuri.def(1)"@/a l2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S l2)a title="append" href="cic:/fakeuri.def(1)"@/a(a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S l1). #S #l1 elim l1 [normalize // | #a #tl #Hind #l2 >a href="cic:/matita/basics/list/reverse_cons.def(7)"reverse_cons/a >a href="cic:/matita/basics/list/reverse_cons.def(7)"reverse_cons/a // qed. -lemma reverse_reverse : ∀S,l. a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S (a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a l. +img class="anchor" src="icons/tick.png" id="reverse_reverse"lemma reverse_reverse : ∀S,l. a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S (a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a l. #S #l elim l // #a #tl #Hind >a href="cic:/matita/basics/list/reverse_cons.def(7)"reverse_cons/a >a href="cic:/matita/basics/list/reverse_append.def(8)"reverse_append/a normalize // qed. (* an elimination principle for lists working on the tail; useful for strings *) -lemma list_elim_left: ∀S.∀P:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S → Prop. P (a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a S) → -(∀a.∀tl.P tl → P (tla title="append" href="cic:/fakeuri.def(1)"@/a(a:a title="cons" href="cic:/fakeuri.def(1)":/a[a title="nil" href="cic:/fakeuri.def(1)"]/a))) → ∀l. P l. +img class="anchor" src="icons/tick.png" id="list_elim_left"lemma list_elim_left: ∀S.∀P:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S → Prop. P (a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a S) → +(∀a.∀tl.P tl → P (tla title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a))) → ∀l. P l. #S #P #Pnil #Pstep #l <(a href="cic:/matita/basics/list/reverse_reverse.def(9)"reverse_reverse/a … l) generalize in match (a href="cic:/matita/basics/list/reverse.def(2)"reverse/a S l); #l elim l // #a #tl #H >a href="cic:/matita/basics/list/reverse_cons.def(7)"reverse_cons/a @Pstep // @@ -166,7 +166,7 @@ qed. (**************************** length *******************************) -let rec length (A:Type[0]) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l ≝ +img class="anchor" src="icons/tick.png" id="length"let rec length (A:Type[0]) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l ≝ match l with [ nil ⇒ a title="natural number" href="cic:/fakeuri.def(1)"0/a | cons a tl ⇒ a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a (length A tl)]. @@ -174,19 +174,19 @@ let rec length (A:Type[0]) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)" notation "|M|" non associative with precedence 60 for @{'norm $M}. interpretation "norm" 'norm l = (length ? l). -lemma length_append: ∀A.∀l1,l2:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A. - |l1a title="append" href="cic:/fakeuri.def(1)"@/al2a title="norm" href="cic:/fakeuri.def(1)"|/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a |l1a title="norm" href="cic:/fakeuri.def(1)"|/aa title="natural plus" href="cic:/fakeuri.def(1)"+/a|l2a title="norm" href="cic:/fakeuri.def(1)"|/a. +img class="anchor" src="icons/tick.png" id="length_append"lemma length_append: ∀A.∀l1,l2:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A. + a title="norm" href="cic:/fakeuri.def(1)"|/al1a title="append" href="cic:/fakeuri.def(1)"@/al2a title="norm" href="cic:/fakeuri.def(1)"|/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="norm" href="cic:/fakeuri.def(1)"|/al1a title="norm" href="cic:/fakeuri.def(1)"|/aa title="natural plus" href="cic:/fakeuri.def(1)"+/aa title="norm" href="cic:/fakeuri.def(1)"|/al2a title="norm" href="cic:/fakeuri.def(1)"|/a. #A #l1 elim l1 // normalize /span class="autotactic"2span class="autotrace" trace /span/span/ qed. -let rec nth n (A:Type[0]) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) (d:A) ≝ +img class="anchor" src="icons/tick.png" id="nth"let rec nth n (A:Type[0]) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) (d:A) ≝ match n with [O ⇒ a href="cic:/matita/basics/list/hd.def(1)"hd/a A l d |S m ⇒ nth m A (a href="cic:/matita/basics/list/tail.def(1)"tail/a A l) d]. (***************************** fold *******************************) -let rec fold (A,B:Type[0]) (op:B → B → B) (b:B) (p:A→a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a) (f:A→B) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l :B ≝ +img class="anchor" src="icons/tick.png" id="fold"let rec fold (A,B:Type[0]) (op:B → B → B) (b:B) (p:A→a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a) (f:A→B) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l :B ≝ match l with [ nil ⇒ b | cons a l ⇒ if (p a) then (op (f a) (fold A B op b p f l)) @@ -203,38 +203,38 @@ for @{'fold $op $nil (λ${ident i}.true) (λ${ident i}. $f) $l}. interpretation "\fold" 'fold op nil p f l = (fold ? ? op nil p f l). -theorem fold_true: +img class="anchor" src="icons/tick.png" id="fold_true"theorem fold_true: ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f:A→B. p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → - \fold[op,nil]_{i ∈ a:a title="cons" href="cic:/fakeuri.def(1)":/al| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a - op (f a) \fold[op,nil]_{i ∈ l| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). + a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a + op (f a) a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ l| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). #A #B #a #l #p #op #nil #f #pa normalize >pa // qed. -theorem fold_false: +img class="anchor" src="icons/tick.png" id="fold_false"theorem fold_false: ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f. -p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → \fold[op,nil]_{i ∈ a:a title="cons" href="cic:/fakeuri.def(1)":/al| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a - \fold[op,nil]_{i ∈ l| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). +p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a + a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ l| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). #A #B #a #l #p #op #nil #f #pa normalize >pa // qed. -theorem fold_filter: +img class="anchor" src="icons/tick.png" id="fold_filter"theorem fold_filter: ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f:A →B. - \fold[op,nil]_{i ∈ l| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a - \fold[op,nil]_{i ∈ (a href="cic:/matita/basics/list/filter.def(2)"filter/a A p l)a title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). + a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ l| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a + a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ (a href="cic:/matita/basics/list/filter.def(2)"filter/a A p l)a title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). #A #B #a #l #p #op #nil #f elim l // #a #tl #Hind cases(a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (p a)) #pa [ >a href="cic:/matita/basics/list/filter_true.def(3)"filter_true/a // > a href="cic:/matita/basics/list/fold_true.def(3)"fold_true/a // >a href="cic:/matita/basics/list/fold_true.def(3)"fold_true/a // | >a href="cic:/matita/basics/list/filter_false.def(3)"filter_false/a // >a href="cic:/matita/basics/list/fold_false.def(3)"fold_false/a // ] qed. -record Aop (A:Type[0]) (nil:A) : Type[0] ≝ +img class="anchor" src="icons/tick.png" id="Aop"record Aop (A:Type[0]) (nil:A) : Type[0] ≝ {op :2> A → A → A; nill:∀a. op nil a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a; nilr:∀a. op a nil a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a; assoc: ∀a,b,c.op a (op b c) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a op (op a b) c }. -theorem fold_sum: ∀A,B. ∀I,J:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.∀nil.∀op:a href="cic:/matita/basics/list/Aop.ind(1,0,2)"Aop/a B nil.∀f. - op (\fold[op,nil]_{i∈Ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i)) (\fold[op,nil]_{i∈Ja title="\fold" href="cic:/fakeuri.def(1)"}/a (f i)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a - \fold[op,nil]_{i∈(Ia title="append" href="cic:/fakeuri.def(1)"@/aJ)a title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). +img class="anchor" src="icons/tick.png" id="fold_sum"theorem fold_sum: ∀A,B. ∀I,J:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.∀nil.∀op:a href="cic:/matita/basics/list/Aop.ind(1,0,2)"Aop/a B nil.∀f. + op (a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i∈Ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i)) (a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i∈Ja title="\fold" href="cic:/fakeuri.def(1)"}/a (f i)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a + a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i∈(Ia title="append" href="cic:/fakeuri.def(1)"@/aJ)a title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). #A #B #I #J #nil #op #f (elim I) normalize [>a href="cic:/matita/basics/list/nill.fix(0,2,2)"nill/a //|#a #tl #Hind <a href="cic:/matita/basics/list/assoc.fix(0,2,2)"assoc/a //] qed. \ No newline at end of file diff --git a/weblib/basics/list2.ma b/weblib/basics/list2.ma index 2543b3084..5d9929fc0 100644 --- a/weblib/basics/list2.ma +++ b/weblib/basics/list2.ma @@ -15,7 +15,7 @@ include "basics/list.ma". include "arithmetics/nat.ma". -let rec length (A:Type[0]) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l ≝ +img class="anchor" src="icons/tick.png" id="length"let rec length (A:Type[0]) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l ≝ match l with [ nil ⇒ a title="natural number" href="cic:/fakeuri.def(1)"0/a | cons a tl ⇒ a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a (length A tl)]. @@ -23,7 +23,7 @@ let rec length (A:Type[0]) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)" notation "|M|" non associative with precedence 60 for @{'norm $M}. interpretation "norm" 'norm l = (length ? l). -let rec nth n (A:Type[0]) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) (d:A) ≝ +img class="anchor" src="icons/tick.png" id="nth"let rec nth n (A:Type[0]) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) (d:A) ≝ match n with [O ⇒ a href="cic:/matita/basics/list/hd.def(1)"hd/a A l d - |S m ⇒ nth m A (a href="cic:/matita/basics/list/tail.def(1)"tail/a A l) d]. + |S m ⇒ nth m A (a href="cic:/matita/basics/list/tail.def(1)"tail/a A l) d]. \ No newline at end of file diff --git a/weblib/basics/logic.ma b/weblib/basics/logic.ma index 527e48429..ffe1498f5 100644 --- a/weblib/basics/logic.ma +++ b/weblib/basics/logic.ma @@ -14,64 +14,64 @@ include "hints_declaration.ma". (* propositional equality *) -inductive eq (A:Type[2]) (x:A) : A → Prop ≝ +img class="anchor" src="icons/tick.png" id="eq"inductive eq (A:Type[2]) (x:A) : A → Prop ≝ refl: eq A x x. interpretation "leibnitz's equality" 'eq t x y = (eq t x y). interpretation "leibniz reflexivity" 'refl = refl. -lemma eq_rect_r: +img class="anchor" src="icons/tick.png" id="eq_rect_r"lemma eq_rect_r: ∀A.∀a,x.∀p:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a.∀P: ∀x:A. a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a → Type[3]. P a (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A a) → P x p. #A #a #x #p (cases p) // qed. -lemma eq_ind_r : +img class="anchor" src="icons/tick.png" id="eq_ind_r"lemma eq_ind_r : ∀A.∀a.∀P: ∀x:A. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a → Prop. P a (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A a) → ∀x.∀p:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a.P x p. #A #a #P #p #x0 #p0; @(a href="cic:/matita/basics/logic/eq_rect_r.def(1)"eq_rect_r/a ? ? ? p0) //; qed. -lemma eq_rect_Type0_r: +img class="anchor" src="icons/tick.png" id="eq_rect_Type0_r"lemma eq_rect_Type0_r: ∀A.∀a.∀P: ∀x:A. a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a → Type[0]. P a (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A a) → ∀x.∀p:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a.P x p. #A #a #P #H #x #p lapply H lapply P cases p; //; qed. -lemma eq_rect_Type1_r: +img class="anchor" src="icons/tick.png" id="eq_rect_Type1_r"lemma eq_rect_Type1_r: ∀A.∀a.∀P: ∀x:A. a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a → Type[1]. P a (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A a) → ∀x.∀p:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a.P x p. #A #a #P #H #x #p lapply H lapply P cases p; //; qed. -lemma eq_rect_Type2_r: +img class="anchor" src="icons/tick.png" id="eq_rect_Type2_r"lemma eq_rect_Type2_r: ∀A.∀a.∀P: ∀x:A. a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a → Type[2]. P a (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A a) → ∀x.∀p:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a.P x p. #A #a #P #H #x #p lapply H lapply P cases p; //; qed. -lemma eq_rect_Type3_r: +img class="anchor" src="icons/tick.png" id="eq_rect_Type3_r"lemma eq_rect_Type3_r: ∀A.∀a.∀P: ∀x:A. a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a → Type[3]. P a (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A a) → ∀x.∀p:a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a ? x a.P x p. #A #a #P #H #x #p lapply H lapply P cases p; //; qed. -theorem rewrite_l: ∀A:Type[2].∀x.∀P:A → Type[2]. P x → ∀y. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y → P y. +img class="anchor" src="icons/tick.png" id="rewrite_l"theorem rewrite_l: ∀A:Type[2].∀x.∀P:A → Type[2]. P x → ∀y. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y → P y. #A #x #P #Hx #y #Heq (cases Heq); //; qed. -theorem sym_eq: ∀A.∀x,y:A. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y → y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x. +img class="anchor" src="icons/tick.png" id="sym_eq"theorem sym_eq: ∀A.∀x,y:A. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y → y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x. #A #x #y #Heq @(a href="cic:/matita/basics/logic/rewrite_l.def(1)"rewrite_l/a A x (λz.za title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax)) // qed. -theorem rewrite_r: ∀A:Type[2].∀x.∀P:A → Type[2]. P x → ∀y. y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x → P y. +img class="anchor" src="icons/tick.png" id="rewrite_r"theorem rewrite_r: ∀A:Type[2].∀x.∀P:A → Type[2]. P x → ∀y. y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x → P y. #A #x #P #Hx #y #Heq (cases (a href="cic:/matita/basics/logic/sym_eq.def(2)"sym_eq/a ? ? ? Heq)); //; qed. -theorem eq_coerc: ∀A,B:Type[0].A→(Aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aB)→B. +img class="anchor" src="icons/tick.png" id="eq_coerc"theorem eq_coerc: ∀A,B:Type[0].A→(Aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aB)→B. #A #B #Ha #Heq (elim Heq); //; qed. -theorem trans_eq : ∀A.∀x,y,z:A. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y → y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a z → x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a z. +img class="anchor" src="icons/tick.png" id="trans_eq"theorem trans_eq : ∀A.∀x,y,z:A. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y → y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a z → x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a z. #A #x #y #z #H1 #H2 >H1; //; qed. -theorem eq_f: ∀A,B.∀f:A→B.∀x,y:A. xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ay → f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f y. +img class="anchor" src="icons/tick.png" id="eq_f"theorem eq_f: ∀A,B.∀f:A→B.∀x,y:A. xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ay → f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f y. #A #B #f #x #y #H >H; //; qed. (* deleterio per auto? *) -theorem eq_f2: ∀A,B,C.∀f:A→B→C. +img class="anchor" src="icons/tick.png" id="eq_f2"theorem eq_f2: ∀A,B,C.∀f:A→B→C. ∀x1,x2:A.∀y1,y2:B. x1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax2 → y1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ay2 → f x1 y1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f x2 y2. #A #B #C #f #x1 #x2 #y1 #y2 #E1 #E2 >E1; >E2; //; qed. -lemma eq_f3: ∀A,B,C,D.∀f:A→B→C->D. +img class="anchor" src="icons/tick.png" id="eq_f3"lemma eq_f3: ∀A,B,C,D.∀f:A→B→C->D. ∀x1,x2:A.∀y1,y2:B. ∀z1,z2:C. x1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax2 → y1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ay2 → z1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/az2 → f x1 y1 z1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f x2 y2 z2. #A #B #C #D #f #x1 #x2 #y1 #y2 #z1 #z2 #E1 #E2 #E3 >E1; >E2; >E3 //; qed. @@ -88,20 +88,20 @@ unification hint 0 ≔ T,a,b; (********** connectives ********) -inductive True: Prop ≝ +img class="anchor" src="icons/tick.png" id="True"inductive True: Prop ≝ I : True. -inductive False: Prop ≝ . +img class="anchor" src="icons/tick.png" id="False"inductive False: Prop ≝ . (* ndefinition Not: Prop → Prop ≝ λA. A → False. *) -inductive Not (A:Prop): Prop ≝ +img class="anchor" src="icons/tick.png" id="Not"inductive Not (A:Prop): Prop ≝ nmk: (A → a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a) → Not A. interpretation "logical not" 'not x = (Not x). -theorem absurd : ∀A:Prop. A → a title="logical not" href="cic:/fakeuri.def(1)"¬/aA → a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a. +img class="anchor" src="icons/tick.png" id="absurd"theorem absurd : ∀A:Prop. A → a title="logical not" href="cic:/fakeuri.def(1)"¬/aA → a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a. #A #H #Hn (elim Hn); /span class="autotactic"2span class="autotrace" trace /span/span/; qed. (* @@ -109,81 +109,81 @@ ntheorem absurd : ∀ A,C:Prop. A → ¬A → C. #A; #C; #H; #Hn; nelim (Hn H). nqed. *) -theorem not_to_not : ∀A,B:Prop. (A → B) → a title="logical not" href="cic:/fakeuri.def(1)"¬/aB →a title="logical not" href="cic:/fakeuri.def(1)"¬/aA. +img class="anchor" src="icons/tick.png" id="not_to_not"theorem not_to_not : ∀A,B:Prop. (A → B) → a title="logical not" href="cic:/fakeuri.def(1)"¬/aB →a title="logical not" href="cic:/fakeuri.def(1)"¬/aA. /span class="autotactic"4span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a, a href="cic:/matita/basics/logic/Not.con(0,1,1)"nmk/a/span/span/; qed. (* inequality *) interpretation "leibnitz's non-equality" 'neq t x y = (Not (eq t x y)). -theorem sym_not_eq: ∀A.∀x,y:A. x a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a y → y a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a x. +img class="anchor" src="icons/tick.png" id="sym_not_eq"theorem sym_not_eq: ∀A.∀x,y:A. x a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a y → y a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a x. /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a, a href="cic:/matita/basics/logic/Not.con(0,1,1)"nmk/a/span/span/; qed. (* and *) -inductive And (A,B:Prop) : Prop ≝ +img class="anchor" src="icons/tick.png" id="And"inductive And (A,B:Prop) : Prop ≝ conj : A → B → And A B. interpretation "logical and" 'and x y = (And x y). -theorem proj1: ∀A,B:Prop. A a title="logical and" href="cic:/fakeuri.def(1)"∧/a B → A. +img class="anchor" src="icons/tick.png" id="proj1"theorem proj1: ∀A,B:Prop. A a title="logical and" href="cic:/fakeuri.def(1)"∧/a B → A. #A #B #AB (elim AB) //; qed. -theorem proj2: ∀ A,B:Prop. A a title="logical and" href="cic:/fakeuri.def(1)"∧/a B → B. +img class="anchor" src="icons/tick.png" id="proj2"theorem proj2: ∀ A,B:Prop. A a title="logical and" href="cic:/fakeuri.def(1)"∧/a B → B. #A #B #AB (elim AB) //; qed. (* or *) -inductive Or (A,B:Prop) : Prop ≝ +img class="anchor" src="icons/tick.png" id="Or"inductive Or (A,B:Prop) : Prop ≝ or_introl : A → (Or A B) | or_intror : B → (Or A B). interpretation "logical or" 'or x y = (Or x y). -definition decidable : Prop → Prop ≝ +img class="anchor" src="icons/tick.png" id="decidable"definition decidable : Prop → Prop ≝ λ A:Prop. A a title="logical or" href="cic:/fakeuri.def(1)"∨/a a title="logical not" href="cic:/fakeuri.def(1)"¬/a A. (* exists *) -inductive ex (A:Type[0]) (P:A → Prop) : Prop ≝ +img class="anchor" src="icons/tick.png" id="ex"inductive ex (A:Type[0]) (P:A → Prop) : Prop ≝ ex_intro: ∀ x:A. P x → ex A P. interpretation "exists" 'exists x = (ex ? x). -inductive ex2 (A:Type[0]) (P,Q:A →Prop) : Prop ≝ +img class="anchor" src="icons/tick.png" id="ex2"inductive ex2 (A:Type[0]) (P,Q:A →Prop) : Prop ≝ ex_intro2: ∀ x:A. P x → Q x → ex2 A P Q. (* iff *) -definition iff := +img class="anchor" src="icons/tick.png" id="iff"definition iff := λ A,B. (A → B) a title="logical and" href="cic:/fakeuri.def(1)"∧/a (B → A). interpretation "iff" 'iff a b = (iff a b). -lemma iff_sym: ∀A,B. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → B a title="iff" href="cic:/fakeuri.def(1)"↔/a A. +img class="anchor" src="icons/tick.png" id="iff_sym"lemma iff_sym: ∀A,B. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → B a title="iff" href="cic:/fakeuri.def(1)"↔/a A. #A #B * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ qed. -lemma iff_trans:∀A,B,C. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → B a title="iff" href="cic:/fakeuri.def(1)"↔/a C → A a title="iff" href="cic:/fakeuri.def(1)"↔/a C. +img class="anchor" src="icons/tick.png" id="iff_trans"lemma iff_trans:∀A,B,C. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → B a title="iff" href="cic:/fakeuri.def(1)"↔/a C → A a title="iff" href="cic:/fakeuri.def(1)"↔/a C. #A #B #C * #H1 #H2 * #H3 #H4 % /span class="autotactic"3span class="autotrace" trace /span/span/ qed. -lemma iff_not: ∀A,B. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → a title="logical not" href="cic:/fakeuri.def(1)"¬/aA a title="iff" href="cic:/fakeuri.def(1)"↔/a a title="logical not" href="cic:/fakeuri.def(1)"¬/aB. +img class="anchor" src="icons/tick.png" id="iff_not"lemma iff_not: ∀A,B. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → a title="logical not" href="cic:/fakeuri.def(1)"¬/aA a title="iff" href="cic:/fakeuri.def(1)"↔/a a title="logical not" href="cic:/fakeuri.def(1)"¬/aB. #A #B * #H1 #H2 % /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a/span/span/ qed. -lemma iff_and_l: ∀A,B,C. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → C a title="logical and" href="cic:/fakeuri.def(1)"∧/a A a title="iff" href="cic:/fakeuri.def(1)"↔/a C a title="logical and" href="cic:/fakeuri.def(1)"∧/a B. +img class="anchor" src="icons/tick.png" id="iff_and_l"lemma iff_and_l: ∀A,B,C. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → C a title="logical and" href="cic:/fakeuri.def(1)"∧/a A a title="iff" href="cic:/fakeuri.def(1)"↔/a C a title="logical and" href="cic:/fakeuri.def(1)"∧/a B. #A #B #C * #H1 #H2 % * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ qed. -lemma iff_and_r: ∀A,B,C. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → A a title="logical and" href="cic:/fakeuri.def(1)"∧/a C a title="iff" href="cic:/fakeuri.def(1)"↔/a B a title="logical and" href="cic:/fakeuri.def(1)"∧/a C. +img class="anchor" src="icons/tick.png" id="iff_and_r"lemma iff_and_r: ∀A,B,C. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → A a title="logical and" href="cic:/fakeuri.def(1)"∧/a C a title="iff" href="cic:/fakeuri.def(1)"↔/a B a title="logical and" href="cic:/fakeuri.def(1)"∧/a C. #A #B #C * #H1 #H2 % * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ qed. -lemma iff_or_l: ∀A,B,C. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → C a title="logical or" href="cic:/fakeuri.def(1)"∨/a A a title="iff" href="cic:/fakeuri.def(1)"↔/a C a title="logical or" href="cic:/fakeuri.def(1)"∨/a B. +img class="anchor" src="icons/tick.png" id="iff_or_l"lemma iff_or_l: ∀A,B,C. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → C a title="logical or" href="cic:/fakeuri.def(1)"∨/a A a title="iff" href="cic:/fakeuri.def(1)"↔/a C a title="logical or" href="cic:/fakeuri.def(1)"∨/a B. #A #B #C * #H1 #H2 % * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ qed. -lemma iff_or_r: ∀A,B,C. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → A a title="logical or" href="cic:/fakeuri.def(1)"∨/a C a title="iff" href="cic:/fakeuri.def(1)"↔/a B a title="logical or" href="cic:/fakeuri.def(1)"∨/a C. +img class="anchor" src="icons/tick.png" id="iff_or_r"lemma iff_or_r: ∀A,B,C. A a title="iff" href="cic:/fakeuri.def(1)"↔/a B → A a title="logical or" href="cic:/fakeuri.def(1)"∨/a C a title="iff" href="cic:/fakeuri.def(1)"↔/a B a title="logical or" href="cic:/fakeuri.def(1)"∨/a C. #A #B #C * #H1 #H2 % * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ qed. (* cose per destruct: da rivedere *) -definition R0 ≝ λT:Type[0].λt:T.t. +img class="anchor" src="icons/tick.png" id="R0"definition R0 ≝ λT:Type[0].λt:T.t. -definition R1 ≝ a href="cic:/matita/basics/logic/eq_rect_Type0.fix(0,5,1)"eq_rect_Type0/a. +img class="anchor" src="icons/tick.png" id="R1"definition R1 ≝ a href="cic:/matita/basics/logic/eq_rect_Type0.fix(0,5,1)"eq_rect_Type0/a. (* used for lambda-delta *) -definition R2 : +img class="anchor" src="icons/tick.png" id="R2"definition R2 : ∀T0:Type[0]. ∀a0:T0. ∀T1:∀x0:T0. a0a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax0 → Type[0]. @@ -201,7 +201,7 @@ definition R2 : @a2 qed. -definition R3 : +img class="anchor" src="icons/tick.png" id="R3"definition R3 : ∀T0:Type[0]. ∀a0:T0. ∀T1:∀x0:T0. a0a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax0 → Type[0]. @@ -224,7 +224,7 @@ definition R3 : @a3 qed. -definition R4 : +img class="anchor" src="icons/tick.png" id="R4"definition R4 : ∀T0:Type[0]. ∀a0:T0. ∀T1:∀x0:T0. a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a T0 a0 x0 → Type[0]. @@ -259,14 +259,14 @@ definition R4 : @a4 qed. -definition eqProp ≝ λA:Prop.a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a A. +img class="anchor" src="icons/tick.png" id="eqProp"definition eqProp ≝ λA:Prop.a href="cic:/matita/basics/logic/eq.ind(1,0,2)"eq/a A. (* Example to avoid indexing and the consequential creation of ill typed terms during paramodulation *) -example lemmaK : ∀A.∀x:A.∀h:xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax. a href="cic:/matita/basics/logic/eqProp.def(1)"eqProp/a ? h (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A x). +img class="anchor" src="icons/tick.png" id="lemmaK"example lemmaK : ∀A.∀x:A.∀h:xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ax. a href="cic:/matita/basics/logic/eqProp.def(1)"eqProp/a ? h (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a A x). #A #x #h @(a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? h: a href="cic:/matita/basics/logic/eqProp.def(1)"eqProp/a ? ? ?). qed. -theorem streicherK : ∀T:Type[2].∀t:T.∀P:t a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a t → Type[3].P (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? t) → ∀p.P p. +img class="anchor" src="icons/tick.png" id="streicherK"theorem streicherK : ∀T:Type[2].∀t:T.∀P:t a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a t → Type[3].P (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? t) → ∀p.P p. #T #t #P #H #p >(a href="cic:/matita/basics/logic/lemmaK.def(2)"lemmaK/a T t p) @H qed. diff --git a/weblib/basics/relations.ma b/weblib/basics/relations.ma index 34d6fe11c..995198aac 100644 --- a/weblib/basics/relations.ma +++ b/weblib/basics/relations.ma @@ -12,78 +12,78 @@ include "basics/logic.ma". (********** relations **********) -definition relation : Type[0] → Type[0] +img class="anchor" src="icons/tick.png" id="relation"definition relation : Type[0] → Type[0] ≝ λA.A→A→Prop. -definition reflexive: ∀A.∀R :a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop +img class="anchor" src="icons/tick.png" id="reflexive"definition reflexive: ∀A.∀R :a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop ≝ λA.λR.∀x:A.R x x. -definition symmetric: ∀A.∀R: a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop +img class="anchor" src="icons/tick.png" id="symmetric"definition symmetric: ∀A.∀R: a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop ≝ λA.λR.∀x,y:A.R x y → R y x. -definition transitive: ∀A.∀R:a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop +img class="anchor" src="icons/tick.png" id="transitive"definition transitive: ∀A.∀R:a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop ≝ λA.λR.∀x,y,z:A.R x y → R y z → R x z. -definition irreflexive: ∀A.∀R:a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop +img class="anchor" src="icons/tick.png" id="irreflexive"definition irreflexive: ∀A.∀R:a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop ≝ λA.λR.∀x:A.a title="logical not" href="cic:/fakeuri.def(1)"¬/a(R x x). -definition cotransitive: ∀A.∀R:a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop +img class="anchor" src="icons/tick.png" id="cotransitive"definition cotransitive: ∀A.∀R:a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop ≝ λA.λR.∀x,y:A.R x y → ∀z:A. R x z a title="logical or" href="cic:/fakeuri.def(1)"∨/a R z y. -definition tight_apart: ∀A.∀eq,ap:a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop +img class="anchor" src="icons/tick.png" id="tight_apart"definition tight_apart: ∀A.∀eq,ap:a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop ≝ λA.λeq,ap.∀x,y:A. (a title="logical not" href="cic:/fakeuri.def(1)"¬/a(ap x y) → eq x y) a title="logical and" href="cic:/fakeuri.def(1)"∧/a (eq x y → a title="logical not" href="cic:/fakeuri.def(1)"¬/a(ap x y)). -definition antisymmetric: ∀A.∀R:a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop +img class="anchor" src="icons/tick.png" id="antisymmetric"definition antisymmetric: ∀A.∀R:a href="cic:/matita/basics/relations/relation.def(1)"relation/a A.Prop ≝ λA.λR.∀x,y:A. R x y → a title="logical not" href="cic:/fakeuri.def(1)"¬/a(R y x). (********** functions **********) -definition compose ≝ +img class="anchor" src="icons/tick.png" id="compose"definition compose ≝ λA,B,C:Type[0].λf:B→C.λg:A→B.λx:A.f (g x). interpretation "function composition" 'compose f g = (compose ? ? ? f g). -definition injective: ∀A,B:Type[0].∀ f:A→B.Prop +img class="anchor" src="icons/tick.png" id="injective"definition injective: ∀A,B:Type[0].∀ f:A→B.Prop ≝ λA,B.λf.∀x,y:A.f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f y → xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/ay. -definition surjective: ∀A,B:Type[0].∀f:A→B.Prop -≝λA,B.λf.∀z:B.a title="exists" href="cic:/fakeuri.def(1)"∃/ax:A.z a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f x. +img class="anchor" src="icons/tick.png" id="surjective"definition surjective: ∀A,B:Type[0].∀f:A→B.Prop +≝λA,B.λf.∀z:B.a title="exists" href="cic:/fakeuri.def(1)"∃/ax:Aa title="exists" href="cic:/fakeuri.def(1)"./az a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f x. -definition commutative: ∀A:Type[0].∀f:A→A→A.Prop +img class="anchor" src="icons/tick.png" id="commutative"definition commutative: ∀A:Type[0].∀f:A→A→A.Prop ≝ λA.λf.∀x,y.f x y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f y x. -definition commutative2: ∀A,B:Type[0].∀f:A→A→B.Prop +img class="anchor" src="icons/tick.png" id="commutative2"definition commutative2: ∀A,B:Type[0].∀f:A→A→B.Prop ≝ λA,B.λf.∀x,y.f x y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f y x. -definition associative: ∀A:Type[0].∀f:A→A→A.Prop +img class="anchor" src="icons/tick.png" id="associative"definition associative: ∀A:Type[0].∀f:A→A→A.Prop ≝ λA.λf.∀x,y,z.f (f x y) z a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a f x (f y z). (* functions and relations *) -definition monotonic : ∀A:Type[0].∀R:A→A→Prop. +img class="anchor" src="icons/tick.png" id="monotonic"definition monotonic : ∀A:Type[0].∀R:A→A→Prop. ∀f:A→A.Prop ≝ λA.λR.λf.∀x,y:A.R x y → R (f x) (f y). (* functions and functions *) -definition distributive: ∀A:Type[0].∀f,g:A→A→A.Prop +img class="anchor" src="icons/tick.png" id="distributive"definition distributive: ∀A:Type[0].∀f,g:A→A→A.Prop ≝ λA.λf,g.∀x,y,z:A. f x (g y z) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a g (f x y) (f x z). -definition distributive2: ∀A,B:Type[0].∀f:A→B→B.∀g:B→B→B.Prop +img class="anchor" src="icons/tick.png" id="distributive2"definition distributive2: ∀A,B:Type[0].∀f:A→B→B.∀g:B→B→B.Prop ≝ λA,B.λf,g.∀x:A.∀y,z:B. f x (g y z) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a g (f x y) (f x z). -lemma injective_compose : ∀A,B,C,f,g. +img class="anchor" src="icons/tick.png" id="injective_compose"lemma injective_compose : ∀A,B,C,f,g. a href="cic:/matita/basics/relations/injective.def(1)"injective/a A B f → a href="cic:/matita/basics/relations/injective.def(1)"injective/a B C g → a href="cic:/matita/basics/relations/injective.def(1)"injective/a A C (λx.g (f x)). -/3/; qed. +/span class="autotactic"3span class="autotrace" trace /span/span/; qed. (* extensional equality *) -definition exteqP: ∀A:Type[0].∀P,Q:A→Prop.Prop ≝ +img class="anchor" src="icons/tick.png" id="exteqP"definition exteqP: ∀A:Type[0].∀P,Q:A→Prop.Prop ≝ λA.λP,Q.∀a.a href="cic:/matita/basics/logic/iff.def(1)"iff/a (P a) (Q a). -definition exteqR: ∀A,B:Type[0].∀R,S:A→B→Prop.Prop ≝ +img class="anchor" src="icons/tick.png" id="exteqR"definition exteqR: ∀A,B:Type[0].∀R,S:A→B→Prop.Prop ≝ λA,B.λR,S.∀a.∀b.a href="cic:/matita/basics/logic/iff.def(1)"iff/a (R a b) (S a b). -definition exteqF: ∀A,B:Type[0].∀f,g:A→B.Prop ≝ +img class="anchor" src="icons/tick.png" id="exteqF"definition exteqF: ∀A,B:Type[0].∀f,g:A→B.Prop ≝ λA,B.λf,g.∀a.f a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a g a. notation " x \eqP y " non associative with precedence 45 @@ -102,4 +102,4 @@ notation " f \eqF g " non associative with precedence 45 for @{'eqF ? ? f g}. interpretation "functional extentional equality" -'eqF A B f g = (exteqF A B f g). +'eqF A B f g = (exteqF A B f g). \ No newline at end of file diff --git a/weblib/basics/types.ma b/weblib/basics/types.ma index ed324db0e..17c211675 100644 --- a/weblib/basics/types.ma +++ b/weblib/basics/types.ma @@ -12,51 +12,51 @@ include "basics/logic.ma". (* void *) -inductive void : Type[0] ≝. +img class="anchor" src="icons/tick.png" id="void"inductive void : Type[0] ≝. (* unit *) -inductive unit : Type[0] ≝ it: unit. +img class="anchor" src="icons/tick.png" id="unit"inductive unit : Type[0] ≝ it: unit. (* sum *) -inductive Sum (A,B:Type[0]) : Type[0] ≝ +img class="anchor" src="icons/tick.png" id="Sum"inductive Sum (A,B:Type[0]) : Type[0] ≝ inl : A → Sum A B | inr : B → Sum A B. interpretation "Disjoint union" 'plus A B = (Sum A B). (* option *) -inductive option (A:Type[0]) : Type[0] ≝ +img class="anchor" src="icons/tick.png" id="option"inductive option (A:Type[0]) : Type[0] ≝ None : option A | Some : A → option A. -definition option_map : ∀A,B:Type[0]. (A → B) → option A → option B ≝ -λA,B,f,o. match o with [ None ⇒ None B | Some a ⇒ Some B (f a) ]. +img class="anchor" src="icons/tick.png" id="option_map"definition option_map : ∀A,B:Type[0]. (A → B) → a href="cic:/matita/basics/types/option.ind(1,0,1)"option/a A → a href="cic:/matita/basics/types/option.ind(1,0,1)"option/a B ≝ +λA,B,f,o. match o with [ None ⇒ a href="cic:/matita/basics/types/option.con(0,1,1)"None/a B | Some a ⇒ a href="cic:/matita/basics/types/option.con(0,2,1)"Some/a B (f a) ]. -lemma option_map_none : ∀A,B,f,x. - option_map A B f x = None B → x = None A. +img class="anchor" src="icons/tick.png" id="option_map_none"lemma option_map_none : ∀A,B,f,x. + a href="cic:/matita/basics/types/option_map.def(1)"option_map/a A B f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/types/option.con(0,1,1)"None/a B → x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/types/option.con(0,1,1)"None/a A. #A #B #f * [ // | #a #E whd in E:(??%?); destruct ] qed. -lemma option_map_some : ∀A,B,f,x,v. - option_map A B f x = Some B v → ∃y. x = Some ? y ∧ f y = v. +img class="anchor" src="icons/tick.png" id="option_map_some"lemma option_map_some : ∀A,B,f,x,v. + a href="cic:/matita/basics/types/option_map.def(1)"option_map/a A B f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/types/option.con(0,2,1)"Some/a B v → a title="exists" href="cic:/fakeuri.def(1)"∃/ay. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/types/option.con(0,2,1)"Some/a ? y a title="logical and" href="cic:/fakeuri.def(1)"∧/a f y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a v. #A #B #f * [ #v normalize #E destruct | #y #v normalize #E %{y} destruct % // ] qed. -definition option_map_def : ∀A,B:Type[0]. (A → B) → B → option A → B ≝ +img class="anchor" src="icons/tick.png" id="option_map_def"definition option_map_def : ∀A,B:Type[0]. (A → B) → B → a href="cic:/matita/basics/types/option.ind(1,0,1)"option/a A → B ≝ λA,B,f,d,o. match o with [ None ⇒ d | Some a ⇒ f a ]. -lemma refute_none_by_refl : ∀A,B:Type[0]. ∀P:A → B. ∀Q:B → Type[0]. ∀x:option A. ∀H:x = None ? → False. - (∀v. x = Some ? v → Q (P v)) → - Q (match x return λy.x = y → ? with [ Some v ⇒ λ_. P v | None ⇒ λE. match H E in False with [ ] ] (refl ? x)). +img class="anchor" src="icons/tick.png" id="refute_none_by_refl"lemma refute_none_by_refl : ∀A,B:Type[0]. ∀P:A → B. ∀Q:B → Type[0]. ∀x:a href="cic:/matita/basics/types/option.ind(1,0,1)"option/a A. ∀H:x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/types/option.con(0,1,1)"None/a ? → a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a. + (∀v. x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/types/option.con(0,2,1)"Some/a ? v → Q (P v)) → + Q (match x return λy.x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y → ? with [ Some v ⇒ λ_. P v | None ⇒ λE. match H E in False with [ ] ] (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ? x)). #A #B #P #Q * -[ #H cases (H (refl ??)) -| #a #H #p normalize @p @refl +[ #H cases (H (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ??)) +| #a #H #p normalize @p @a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a ] qed. (* sigma *) -record Sig (A:Type[0]) (f:A→Type[0]) : Type[0] ≝ { +img class="anchor" src="icons/tick.png" id="Sig"record Sig (A:Type[0]) (f:A→Type[0]) : Type[0] ≝ { pi1: A ; pi2: f pi1 }. @@ -70,7 +70,7 @@ interpretation "mk_Sig" 'dp x y = (mk_Sig ?? x y). (* Prod *) -record Prod (A,B:Type[0]) : Type[0] ≝ { +img class="anchor" src="icons/tick.png" id="Prod"record Prod (A,B:Type[0]) : Type[0] ≝ { fst: A ; snd: B }. @@ -99,94 +99,94 @@ with precedence 90 for @{ 'quadruple $a $b $c $d}. interpretation "Quadruple construction" 'quadruple w x y z = (mk_Prod ? ? (mk_Prod ? ? w x) (mk_Prod ? ? y z)). -theorem eq_pair_fst_snd: ∀A,B.∀p:A × B. - p = 〈 \fst p, \snd p 〉. +img class="anchor" src="icons/tick.png" id="eq_pair_fst_snd"theorem eq_pair_fst_snd: ∀A,B.∀p:A a title="Product" href="cic:/fakeuri.def(1)"×/a B. + p a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p, a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. #A #B #p (cases p) // qed. -lemma fst_eq : ∀A,B.∀a:A.∀b:B. \fst 〈a,b〉 = a. +img class="anchor" src="icons/tick.png" id="fst_eq"lemma fst_eq : ∀A,B.∀a:A.∀b:B. a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a. // qed. -lemma snd_eq : ∀A,B.∀a:A.∀b:B. \snd 〈a,b〉 = b. +img class="anchor" src="icons/tick.png" id="snd_eq"lemma snd_eq : ∀A,B.∀a:A.∀b:B. a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b. // qed. -notation > "hvbox('let' 〈ident x,ident y〉 ≝ t 'in' s)" +notation > "hvbox('let' 〈ident x,ident y〉 ≝ t 'in' s)" with precedence 10 for @{ match $t with [ mk_Prod ${ident x} ${ident y} ⇒ $s ] }. -notation < "hvbox('let' \nbsp hvbox(〈ident x,ident y〉 \nbsp≝ break t \nbsp 'in' \nbsp) break s)" +notation < "hvbox('let' \nbsp hvbox(〈ident x,ident y〉 \nbsp≝ break t \nbsp 'in' \nbsp) break s)" with precedence 10 for @{ match $t with [ mk_Prod (${ident x}:$X) (${ident y}:$Y) ⇒ $s ] }. (* Also extracts an equality proof (useful when not using Russell). *) -notation > "hvbox('let' 〈ident x,ident y〉 'as' ident E ≝ t 'in' s)" +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_Prod ${ident x} ${ident y} ⇒ λ${ident E}.$s ] (refl ? $t) }. (* Prop sigma *) -record PSig (A:Type[0]) (P:A→Prop) : Type[0] ≝ +img class="anchor" src="icons/tick.png" id="PSig"record PSig (A:Type[0]) (P:A→Prop) : Type[0] ≝ {elem:>A; eproof: P elem}. interpretation "subset type" 'sigma x = (PSig ? x). -notation < "hvbox('let' \nbsp hvbox(〈ident x,ident y〉 \nbsp 'as'\nbsp ident E\nbsp ≝ break t \nbsp 'in' \nbsp) break s)" +notation < "hvbox('let' \nbsp hvbox(〈ident x,ident y〉 \nbsp 'as'\nbsp ident E\nbsp ≝ break t \nbsp 'in' \nbsp) break s)" with precedence 10 for @{ match $t return λ${ident k}:$X.$eq $T $k $t → ? with [ mk_Prod (${ident x}:$U) (${ident y}:$W) ⇒ λ${ident E}:$e.$s ] ($refl $T $t) }. -notation > "hvbox('let' 〈ident x,ident y,ident z〉 'as' ident E ≝ t 'in' s)" +notation > "hvbox('let' 〈ident x,ident y,ident z〉 'as' ident E ≝ t 'in' s)" with precedence 10 for @{ match $t return λx.x = $t → ? with [ mk_Prod ${fresh xy} ${ident z} ⇒ match ${fresh xy} return λx. ? = $t → ? with [ mk_Prod ${ident x} ${ident y} ⇒ λ${ident E}.$s ] ] (refl ? $t) }. -notation < "hvbox('let' \nbsp hvbox(〈ident x,ident y,ident z〉 \nbsp'as'\nbsp ident E\nbsp ≝ break t \nbsp 'in' \nbsp) break s)" +notation < "hvbox('let' \nbsp hvbox(〈ident x,ident y,ident z〉 \nbsp'as'\nbsp ident E\nbsp ≝ break t \nbsp 'in' \nbsp) break s)" with precedence 10 for @{ match $t return λ${ident x}.$eq $T $x $t → $U with [ mk_Prod (${fresh xy}:$V) (${ident z}:$Z) ⇒ match ${fresh xy} return λ${ident y}. $eq $R $r $t → ? with [ mk_Prod (${ident x}:$L) (${ident y}:$I) ⇒ λ${ident E}:$J.$s ] ] ($refl $A $t) }. -notation > "hvbox('let' 〈ident w,ident x,ident y,ident z〉 ≝ t 'in' s)" +notation > "hvbox('let' 〈ident w,ident x,ident y,ident z〉 ≝ t 'in' s)" with precedence 10 for @{ match $t with [ mk_Prod ${fresh wx} ${fresh yz} ⇒ match ${fresh wx} with [ mk_Prod ${ident w} ${ident x} ⇒ match ${fresh yz} with [ mk_Prod ${ident y} ${ident z} ⇒ $s ] ] ] }. -notation > "hvbox('let' 〈ident x,ident y,ident z〉 ≝ t 'in' s)" +notation > "hvbox('let' 〈ident x,ident y,ident z〉 ≝ t 'in' s)" with precedence 10 for @{ match $t with [ mk_Prod ${fresh xy} ${ident z} ⇒ match ${fresh xy} with [ mk_Prod ${ident x} ${ident y} ⇒ $s ] ] }. (* This appears to upset automation (previously provable results require greater depth or just don't work), so use example rather than lemma to prevent it being indexed. *) -example contract_pair : ∀A,B.∀e:A×B. (let 〈a,b〉 ≝ e in 〈a,b〉) = e. +img class="anchor" src="icons/tick.png" id="contract_pair"example contract_pair : ∀A,B.∀e:Aa title="Product" href="cic:/fakeuri.def(1)"×/aB. (let 〈a,b〉 ≝ e in a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a e. #A #B * // qed. -lemma extract_pair : ∀A,B,C,D. ∀u:A×B. ∀Q:A → B → C×D. ∀x,y. -((let 〈a,b〉 ≝ u in Q a b) = 〈x,y〉) → -∃a,b. 〈a,b〉 = u ∧ Q a b = 〈x,y〉. -#A #B #C #D * #a #b #Q #x #y normalize #E1 %{a} %{b} % try @refl @E1 qed. +img class="anchor" src="icons/tick.png" id="extract_pair"lemma extract_pair : ∀A,B,C,D. ∀u:Aa title="Product" href="cic:/fakeuri.def(1)"×/aB. ∀Q:A → B → Ca title="Product" href="cic:/fakeuri.def(1)"×/aD. ∀x,y. +((let 〈a,b〉 ≝ u in Q a b) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ax,ya title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) → +a title="exists" href="cic:/fakeuri.def(1)"∃/aa,ba title="exists" href="cic:/fakeuri.def(1)"./a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a u a title="logical and" href="cic:/fakeuri.def(1)"∧/a Q a b a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ax,ya title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. +#A #B #C #D * #a #b #Q #x #y normalize #E1 %{a} %{b} % try @a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a @E1 qed. -lemma breakup_pair : ∀A,B,C:Type[0].∀x. ∀R:C → Prop. ∀P:A → B → C. - R (P (\fst x) (\snd x)) → R (let 〈a,b〉 ≝ x in P a b). -#A #B #C *; normalize /2/ +img class="anchor" src="icons/tick.png" id="breakup_pair"lemma breakup_pair : ∀A,B,C:Type[0].∀x. ∀R:C → Prop. ∀P:A → B → C. + R (P (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a x) (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a x)) → R (let 〈a,b〉 ≝ x in P a b). +#A #B #C *; normalize /span class="autotactic"2span class="autotrace" trace /span/span/ qed. -lemma pair_elim: +img class="anchor" src="icons/tick.png" id="pair_elim"lemma pair_elim: ∀A,B,C: Type[0]. ∀T: A → B → C. ∀p. - ∀P: A×B → C → Prop. - (∀lft, rgt. p = 〈lft,rgt〉 → P 〈lft,rgt〉 (T lft rgt)) → - P p (let 〈lft, rgt〉 ≝ p in T lft rgt). - #A #B #C #T * /2/ + ∀P: Aa title="Product" href="cic:/fakeuri.def(1)"×/aB → C → Prop. + (∀lft, rgt. p a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/alft,rgta title="Pair construction" href="cic:/fakeuri.def(1)"〉/a → P a title="Pair construction" href="cic:/fakeuri.def(1)"〈/alft,rgta title="Pair construction" href="cic:/fakeuri.def(1)"〉/a (T lft rgt)) → + P p (let 〈lft, rgt〉 ≝ p in T lft rgt). + #A #B #C #T * /span class="autotactic"2span class="autotrace" trace /span/span/ qed. -lemma pair_elim2: +img class="anchor" src="icons/tick.png" id="pair_elim2"lemma pair_elim2: ∀A,B,C,C': Type[0]. ∀T: A → B → C. ∀T': A → B → C'. ∀p. - ∀P: A×B → C → C' → Prop. - (∀lft, rgt. p = 〈lft,rgt〉 → P 〈lft,rgt〉 (T lft rgt) (T' lft rgt)) → - P p (let 〈lft, rgt〉 ≝ p in T lft rgt) (let 〈lft, rgt〉 ≝ p in T' lft rgt). - #A #B #C #C' #T #T' * /2/ -qed. + ∀P: Aa title="Product" href="cic:/fakeuri.def(1)"×/aB → C → C' → Prop. + (∀lft, rgt. p a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/alft,rgta title="Pair construction" href="cic:/fakeuri.def(1)"〉/a → P a title="Pair construction" href="cic:/fakeuri.def(1)"〈/alft,rgta title="Pair construction" href="cic:/fakeuri.def(1)"〉/a (T lft rgt) (T' lft rgt)) → + P p (let 〈lft, rgt〉 ≝ p in T lft rgt) (let 〈lft, rgt〉 ≝ p in T' lft rgt). + #A #B #C #C' #T #T' * /span class="autotactic"2span class="autotrace" trace /span/span/ +qed. \ No newline at end of file diff --git a/weblib/hints_declaration.ma b/weblib/hints_declaration.ma index d04f453f2..588741fa3 100644 --- a/weblib/hints_declaration.ma +++ b/weblib/hints_declaration.ma @@ -58,12 +58,12 @@ notation > "≔ (list0 ( (list1 (ident x) sep , ) opt (: T) ) sep ,) opt (; (lis include "basics/pts.ma". -definition hint_declaration_Type0 ≝ λA:Type[0] .λa,b:A.Prop. -definition hint_declaration_Type1 ≝ λA:Type[1].λa,b:A.Prop. -definition hint_declaration_Type2 ≝ λa,b:Type[2].Prop. -definition hint_declaration_CProp0 ≝ λA:CProp[0].λa,b:A.Prop. -definition hint_declaration_CProp1 ≝ λA:CProp[1].λa,b:A.Prop. -definition hint_declaration_CProp2 ≝ λa,b:CProp[2].Prop. +img class="anchor" src="icons/tick.png" id="hint_declaration_Type0"definition hint_declaration_Type0 ≝ λA:Type[0] .λa,b:A.Prop. +img class="anchor" src="icons/tick.png" id="hint_declaration_Type1"definition hint_declaration_Type1 ≝ λA:Type[1].λa,b:A.Prop. +img class="anchor" src="icons/tick.png" id="hint_declaration_Type2"definition hint_declaration_Type2 ≝ λa,b:Type[2].Prop. +img class="anchor" src="icons/tick.png" id="hint_declaration_CProp0"definition hint_declaration_CProp0 ≝ λA:CProp[0].λa,b:A.Prop. +img class="anchor" src="icons/tick.png" id="hint_declaration_CProp1"definition hint_declaration_CProp1 ≝ λA:CProp[1].λa,b:A.Prop. +img class="anchor" src="icons/tick.png" id="hint_declaration_CProp2"definition hint_declaration_CProp2 ≝ λa,b:CProp[2].Prop. interpretation "hint_decl_Type2" 'hint_decl a b = (hint_declaration_Type2 a b). interpretation "hint_decl_CProp2" 'hint_decl a b = (hint_declaration_CProp2 a b). diff --git a/weblib/test.ma b/weblib/test.ma index ab54bb5d1..ce8cc2355 100644 --- a/weblib/test.ma +++ b/weblib/test.ma @@ -1,8 +1,42 @@ -inductive nat : Set ≝ -| O : nat -| S : nat span class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"span class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"span class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"→/span/span/span nat. +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +(* + * QUICK START GUIDE + * ================= + * MatitaWeb is a web interface for the Matita interactive theorem prover. + * Its user interface is composed of two main panes: + * + * -> the script area, whose content you are currently reading, which is + * used to input definitions and theorems to be checked by Matita; + * -> a top area with a title bar and a toolbar providing all the + * interactive operations apart from proof authoring. + * + * The toolbar on the top is divided in two sections: + * -> the large blue tiles provide the basic operations concerning the + * execution and checking of a proof script, including step-by-step + * execution + * -> the smaller circular buttons on the right provide script management + * operations (create, open, save...) + * Hover the mouse pointer over these elements to show tooltips explaining + * their purpose. + * + * If you are new to Matita and/or interactive theorem proving, we have a + * tutorial in 10 parts for you. To load it: + * 1) click on the "Open script" button in the toolbar (the second circular + * button, whose icon shows a folder) + * 2) in the dialog box, select the "tutorial/chapter1.ma" script, then + * press OK (scripts chapter2.ma ... chapter10.ma contain more advanced + * examples that the user is encouraged to try after chapter1.ma) + * 3) follow the instructions in the script. + * + *) -let rec plus n m ≝ -match n with -[O ⇒ m -|(S p) ⇒ S (plus p m)]. \ No newline at end of file diff --git a/weblib/tutorial/chapter1.ma b/weblib/tutorial/chapter1.ma index ef764ceae..4cf145d60 100644 --- a/weblib/tutorial/chapter1.ma +++ b/weblib/tutorial/chapter1.ma @@ -44,24 +44,24 @@ few preliminary notions not worth discussing for the moment. include "basics/logic.ma". -inductive bank: Type[0] ≝ +img class="anchor" src="icons/tick.png" id="bank"inductive bank: Type[0] ≝ | east : bank | west : bank. (* We can now define a simple function computing, for each bank of the river, the opposite one. *) -definition opposite ≝ λs. +img class="anchor" src="icons/tick.png" id="opposite"definition opposite ≝ λs. match s with - [ east ⇒ west - | west ⇒ east + [ east ⇒ a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a + | west ⇒ a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a ]. (* Functions are live entities, and can be actually computed. To check this, let us state the property that the opposite bank of east is west; every lemma needs a name for further reference, and we call it "east_to_west". *) -lemma east_to_west : opposite east = west. +img class="anchor" src="icons/tick.png" id="east_to_west"lemma east_to_west : a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a. (* h2 class="section"The goal window/h2 @@ -102,14 +102,14 @@ lemma performing some book-keeping operations. In this case, we avoid the unnecessary simplification step: // will take care of it. *) -lemma west_to_east : opposite west = east. +img class="anchor" src="icons/tick.png" id="west_to_east"lemma west_to_east : a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a. // qed. (* h2 class="section"Introduction/h2 A slightly more complex problem consists in proving that opposite is idempotent *) -lemma idempotent_opposite : ∀x. opposite (opposite x) = x. +img class="anchor" src="icons/tick.png" id="idempotent_opposite"lemma idempotent_opposite : ∀x. a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a (a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a x) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x. (* we start the proof moving x from the conclusion into the context, that is a (backward) introduction step. Matita syntax for an introduction step is simply @@ -155,9 +155,9 @@ other. Only two cases are possible, leading naturally to the following definition: *) -inductive opp : bank → bank → Prop ≝ -| east_west : opp east west -| west_east : opp west east. +img class="anchor" src="icons/tick.png" id="opp"inductive opp : a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a → a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a → Prop ≝ +| east_west : opp a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a +| west_east : opp a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a. (* In precisely the same way as "bank" is the smallest type containing east and west, opp is the smallest predicate containing the two sub-cases east_west and @@ -169,7 +169,7 @@ Between opp and opposite we have the following relation: opp a b iff a = opposite b Let us prove it, starting from the left to right implication, first *) -lemma opp_to_opposite: ∀a,b. opp a b → a = opposite b. +img class="anchor" src="icons/tick.png" id="opp_to_opposite"lemma opp_to_opposite: ∀a,b. a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a a b → a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a b. (* We start the proof introducing a, b and the hypothesis opp a b, that we call oppab. *) @@ -186,7 +186,7 @@ cases oppab // qed. h2 class="section"Rewriting/h2 Let us come to the opposite direction. *) -lemma opposite_to_opp: ∀a,b. a = opposite b → opp a b. +img class="anchor" src="icons/tick.png" id="opposite_to_opp"lemma opposite_to_opp: ∀a,b. a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter1/opposite.def(1)"opposite/a b → a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a a b. (* As usual, we start introducing a, b and the hypothesis (a = opposite b), that we call eqa. *) @@ -211,44 +211,44 @@ wolf, the cabbage, and the boat. The simplest way to declare such a data type is to use a record. *) -record state : Type[0] ≝ - {goat_pos : bank; - wolf_pos : bank; - cabbage_pos: bank; - boat_pos : bank}. +img class="anchor" src="icons/tick.png" id="state"record state : Type[0] ≝ + {goat_pos : a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a; + wolf_pos : a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a; + cabbage_pos: a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a; + boat_pos : a href="cic:/matita/tutorial/chapter1/bank.ind(1,0,0)"bank/a}. (* When you define a record named foo, the system automatically defines a record constructor named mk_foo. To construct a new record you pass as arguments to mk_foo the values of the record fields *) -definition start ≝ mk_state east east east east. -definition end ≝ mk_state west west west west. +img class="anchor" src="icons/tick.png" id="start"definition start ≝ a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a. +img class="anchor" src="icons/tick.png" id="end"definition end ≝ a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a. (* We must now define the possible moves. A natural way to do it is in the form of a relation (a binary predicate) over states. *) -inductive move : state → state → Prop ≝ -| move_goat: ∀g,g1,w,c. opp g g1 → move (mk_state g w c g) (mk_state g1 w c g1) +img class="anchor" src="icons/tick.png" id="move"inductive move : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → Prop ≝ +| move_goat: ∀g,g1,w,c. a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a g g1 → move (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g w c g) (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g1 w c g1) (* We can move the goat from a bank g to the opposite bank g1 if and only if the boat is on the same bank g of the goat and we move the boat along with it. *) -| move_wolf: ∀g,w,w1,c. opp w w1 → move (mk_state g w c w) (mk_state g w1 c w1) -| move_cabbage: ∀g,w,c,c1.opp c c1 → move (mk_state g w c c) (mk_state g w c1 c1) -| move_boat: ∀g,w,c,b,b1. opp b b1 → move (mk_state g w c b) (mk_state g w c b1). +| move_wolf: ∀g,w,w1,c. a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a w w1 → move (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g w c w) (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g w1 c w1) +| move_cabbage: ∀g,w,c,c1.a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a c c1 → move (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g w c c) (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g w c1 c1) +| move_boat: ∀g,w,c,b,b1. a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a b b1 → move (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g w c b) (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g w c b1). (* A state is safe if either the goat is on the same bank of the boat, or both the wolf and the cabbage are on the opposite bank of the goat. *) -inductive safe_state : state → Prop ≝ -| with_boat : ∀g,w,c.safe_state (mk_state g w c g) -| opposite_side : ∀g,g1,b.opp g g1 → safe_state (mk_state g g1 g1 b). +img class="anchor" src="icons/tick.png" id="safe_state"inductive safe_state : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → Prop ≝ +| with_boat : ∀g,w,c.safe_state (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g w c g) +| opposite_side : ∀g,g1,b.a href="cic:/matita/tutorial/chapter1/opp.ind(1,0,0)"opp/a g g1 → safe_state (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a g g1 g1 b). (* Finally, a state y is reachable from x if either there is a single move leading from x to y, or there is a safe state z such that z is reachable from x and there is a move leading from z to y *) -inductive reachable : state → state → Prop ≝ -| one : ∀x,y.move x y → reachable x y -| more : ∀x,z,y. span style="text-decoration: underline;"/spanreachable x z → safe_state z → span style="text-decoration: underline;"/spanmove z y → reachable x y. +img class="anchor" src="icons/tick.png" id="reachable"inductive reachable : a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → a href="cic:/matita/tutorial/chapter1/state.ind(1,0,0)"state/a → Prop ≝ +| one : ∀x,y.a href="cic:/matita/tutorial/chapter1/move.ind(1,0,0)"move/a x y → reachable x y +| more : ∀x,z,y. span style="text-decoration: underline;"/spanreachable x z → a href="cic:/matita/tutorial/chapter1/safe_state.ind(1,0,0)"safe_state/a z → span style="text-decoration: underline;"/spana href="cic:/matita/tutorial/chapter1/move.ind(1,0,0)"move/a z y → reachable x y. (* h2 class="section"Automation/h2 @@ -260,8 +260,8 @@ applicative nodes). In this case, there is a solution in six moves, and we need a few more applications to handle reachability, and side conditions. The magic number to let automation work is, in this case, 9. *) -lemma problem: reachable start end. -normalize /span class="autotactic"9span class="autotrace" trace one, more, with_boat, opposite_side, move_goat, move_wolf, move_cabbage, move_boat, east_west, west_east/span/span/ qed. +img class="anchor" src="icons/tick.png" id="problem"lemma problem: a href="cic:/matita/tutorial/chapter1/reachable.ind(1,0,0)"reachable/a a href="cic:/matita/tutorial/chapter1/start.def(1)"start/a a href="cic:/matita/tutorial/chapter1/end.def(1)"end/a. +normalize /span class="autotactic"9span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/reachable.con(0,1,0)"one/a, a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a, a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"with_boat/a, a href="cic:/matita/tutorial/chapter1/safe_state.con(0,2,0)"opposite_side/a, a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"move_goat/a, a href="cic:/matita/tutorial/chapter1/move.con(0,2,0)"move_wolf/a, a href="cic:/matita/tutorial/chapter1/move.con(0,3,0)"move_cabbage/a, a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"move_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ qed. (* h2 class="section"Application/h2 @@ -275,8 +275,8 @@ the way it is actually used (e.g. for introduction, rewriting, in an applicative step, and so on). *) -lemma problem1: reachable start end. -normalize @more +img class="anchor" src="icons/tick.png" id="problem1"lemma problem1: a href="cic:/matita/tutorial/chapter1/reachable.ind(1,0,0)"reachable/a a href="cic:/matita/tutorial/chapter1/start.def(1)"start/a a href="cic:/matita/tutorial/chapter1/end.def(1)"end/a. +normalize @a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a (* h2 class="section"Focusing/h2 @@ -314,7 +314,7 @@ We can now proceed in several possible ways. The most straightforward way is to provide the intermediate state, that is [east,west,west,east]. We can do it, by just applying this term. *) - @(mk_state east west west east) + @(a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a) (* This application closes the goal; at present, no goal has the focus on. In order to act on the next goal, we must focus on it using the "|" operator. In @@ -331,7 +331,7 @@ requires /2/ since move_goat opens an additional subgoal. By applying "]" we refocus on the skipped goal, going back to a situation similar to the one we started with. *) - | /span class="autotactic"2span class="autotrace" trace move_goat, east_west/span/span/ ] + | /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"move_goat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a/span/span/ ] (* h2 class="section"Implicit arguments/h2 @@ -342,20 +342,20 @@ already instatated on the next intermediate state. As first argument, we type a question mark that stands for an implicit argument to be guessed by the system. *) -@(more ? (mk_state east west west west)) +@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a ? (a href="cic:/matita/tutorial/chapter1/state.con(0,1,0)"mk_state/a a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a)) (* We now get three independent subgoals, all actives, and two of them are trivial. Wespan style="font-family: Verdana,sans-serif;" /spancan just apply automation to all of them, and it will close the two trivial goals. *) -/span class="autotactic"2span class="autotrace" trace opposite_side, move_boat, east_west, west_east/span/span/ +/span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/safe_state.con(0,2,0)"opposite_side/a, a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"move_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ (* Let us come to the next step, that consists in moving the wolf. Suppose that instead of specifying the next intermediate state, we prefer to specify the next move. In the spirit of the previous example, we can do it in the following way *) -@(more … (move_wolf … )) +@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a … (a href="cic:/matita/tutorial/chapter1/move.con(0,2,0)"move_wolf/a … )) (* The dots stand here for an arbitrary number of implicit arguments, to be guessed by the system. @@ -367,12 +367,12 @@ be arbitrary). The simplest way to proceed is to focus on the bank, that is the fourth subgoal, and explicitly instatiate it. Instead of repeatedly using "|", we can perform focusing by typing "4:" as described by the following command. *) -[4: @east] /span class="autotactic"2span class="autotrace" trace with_boat, east_west/span/span/ +[4: @a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a] /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"with_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a/span/span/ (* Alternatively, we can directly instantiate the bank into the move. Let us complete the proof in this, very readable way. *) -@(more … (move_goat west … )) /span class="autotactic"2span class="autotrace" trace with_boat, west_east/span/span/ -@(more … (move_cabbage ?? east … )) /span class="autotactic"2span class="autotrace" trace opposite_side, east_west, west_east/span/span/ -@(more … (move_boat ??? west … )) /span class="autotactic"2span class="autotrace" trace with_boat, west_east/span/span/ -@one /span class="autotactic"2span class="autotrace" trace move_goat, east_west/span/span/ qed. \ No newline at end of file +@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a … (a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"move_goat/a a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a … )) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"with_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ +@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a … (a href="cic:/matita/tutorial/chapter1/move.con(0,3,0)"move_cabbage/a ?? a href="cic:/matita/tutorial/chapter1/bank.con(0,1,0)"east/a … )) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/safe_state.con(0,2,0)"opposite_side/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ +@(a href="cic:/matita/tutorial/chapter1/reachable.con(0,2,0)"more/a … (a href="cic:/matita/tutorial/chapter1/move.con(0,4,0)"move_boat/a ??? a href="cic:/matita/tutorial/chapter1/bank.con(0,2,0)"west/a … )) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/safe_state.con(0,1,0)"with_boat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,2,0)"west_east/a/span/span/ +@a href="cic:/matita/tutorial/chapter1/reachable.con(0,1,0)"one/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter1/move.con(0,1,0)"move_goat/a, a href="cic:/matita/tutorial/chapter1/opp.con(0,1,0)"east_west/a/span/span/ qed. \ No newline at end of file diff --git a/weblib/tutorial/chapter10.ma b/weblib/tutorial/chapter10.ma index 8fe459cd2..15961e536 100644 --- a/weblib/tutorial/chapter10.ma +++ b/weblib/tutorial/chapter10.ma @@ -6,15 +6,15 @@ include "tutorial/chapter9.ma". (* We say that two pres 〈i_1,b_1〉 and 〈i_1,b_1〉 are {\em cofinal} if and only if b_1 = b_2. *) -definition cofinal ≝ λS.λp:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). +img class="anchor" src="icons/tick.png" id="cofinal"definition cofinal ≝ λS.λp:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p). (* As a corollary of decidable_sem, we have that two expressions e1 and e2 are equivalent iff for any word w the states reachable through w are cofinal. *) -theorem equiv_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2} a title="iff" href="cic:/fakeuri.def(1)"↔/a ∀w.a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e2〉. +img class="anchor" src="icons/tick.png" id="equiv_sem"theorem equiv_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1a title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2a title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="iff" href="cic:/fakeuri.def(1)"↔/a ∀w.a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. #S #e1 #e2 % [#same_sem #w cut (∀b1,b2. a href="cic:/matita/basics/logic/iff.def(1)"iff/a (b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) (b2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) → (b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b2)) @@ -29,12 +29,12 @@ length of w; moreover, so far, we made no assumption over the cardinality of S. Instead of requiring S to be finite, we may restrict the analysis to characters occurring in the given pres. *) -definition occ ≝ λS.λe1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1|)) (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2|)). +img class="anchor" src="icons/tick.png" id="occ"definition occ ≝ λS.λe1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1a title="forget" href="cic:/fakeuri.def(1)"|/a)) (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2a title="forget" href="cic:/fakeuri.def(1)"|/a)). -lemma occ_enough: ∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. -(∀w.(a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S w (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a S e1 e2))→ a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e2〉) - →∀w.a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e2〉. +img class="anchor" src="icons/tick.png" id="occ_enough"lemma occ_enough: ∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. +(∀w.(a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S w (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a S e1 e2))→ a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) + →∀w.a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. #S #e1 #e2 #H #w cases (a href="cic:/matita/tutorial/chapter5/decidable_sublist.def(6)"decidable_sublist/a S w (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a S e1 e2)) [@H] -H #H >a href="cic:/matita/tutorial/chapter9/to_pit.def(10)"to_pit/a [2: @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 #a #memba @a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l1.def(6)"sublist_unique_append_l1/a @H1 //] @@ -45,9 +45,9 @@ qed. (* The following is a stronger version of equiv_sem, relative to characters occurring the given regular expressions. *) -lemma equiv_sem_occ: ∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. -(∀w.(a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S w (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a S e1 e2))→ a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e2〉) -→ a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1}a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2}. +img class="anchor" src="icons/tick.png" id="equiv_sem_occ"lemma equiv_sem_occ: ∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. +(∀w.(a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S w (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a S e1 e2))→ a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) +→ a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1a title="in_prl" href="cic:/fakeuri.def(1)"}/aa title="extensional equality" href="cic:/fakeuri.def(1)"=/a1a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2a title="in_prl" href="cic:/fakeuri.def(1)"}/a. #S #e1 #e2 #H @(a href="cic:/matita/basics/logic/proj2.def(2)"proj2/a … (a href="cic:/matita/tutorial/chapter10/equiv_sem.def(16)"equiv_sem/a …)) @a href="cic:/matita/tutorial/chapter10/occ_enough.def(11)"occ_enough/a #w @H qed. @@ -58,10 +58,10 @@ We say that a list of pairs of pres is a bisimulation if it is closed w.r.t. moves, and all its members are cofinal. *) -definition sons ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λl:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S.λp:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). - a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a ?? (λa.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p)),a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p))〉) l. +img class="anchor" src="icons/tick.png" id="sons"definition sons ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λl:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S.λp:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). + a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a ?? (λa.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p)),a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p))a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) l. -lemma memb_sons: ∀S,l.∀p,q:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? p (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a ? l q) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → +img class="anchor" src="icons/tick.png" id="memb_sons"lemma memb_sons: ∀S,l.∀p,q:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? p (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a ? l q) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a title="exists" href="cic:/fakeuri.def(1)"∃/aa.(a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a q)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p a title="logical and" href="cic:/fakeuri.def(1)"∧/a a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a q)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p). #S #l elim l [#p #q normalize in ⊢ (%→?); #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/] @@ -69,15 +69,15 @@ lemma memb_sons: ∀S,l.∀p,q:(a href="cic:/matita/tutorial/chapter7/pre.def(1 [#H @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … a) >(\P H) /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ |#H @Hind @H] qed. -definition is_bisim ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λl:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?.λalpha:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S. +img class="anchor" src="icons/tick.png" id="is_bisim"definition is_bisim ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λl:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?.λalpha:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S. ∀p:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? p l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter10/cofinal.def(2)"cofinal/a ? p a title="logical and" href="cic:/fakeuri.def(1)"∧/a (a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a ? (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a ? alpha p) l). (* Using lemma equiv_sem_occ it is easy to prove the following result: *) -lemma bisim_to_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?.∀e1,e2: a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a href="cic:/matita/tutorial/chapter10/is_bisim.def(8)"is_bisim/a S l (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a S e1 e2) → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ?a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ae1,e2〉 l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1}a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2}. +img class="anchor" src="icons/tick.png" id="bisim_to_sem"lemma bisim_to_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?.∀e1,e2: a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a href="cic:/matita/tutorial/chapter10/is_bisim.def(8)"is_bisim/a S l (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a S e1 e2) → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ae1,e2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1a title="in_prl" href="cic:/fakeuri.def(1)"}/aa title="extensional equality" href="cic:/fakeuri.def(1)"=/a1a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2a title="in_prl" href="cic:/fakeuri.def(1)"}/a. #S #l #e1 #e2 #Hbisim #Hmemb @a href="cic:/matita/tutorial/chapter10/equiv_sem_occ.def(17)"equiv_sem_occ/a -#w #Hsub @(a href="cic:/matita/basics/logic/proj1.def(2)"proj1/a … (Hbisim a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e2〉 ?)) +#w #Hsub @(a href="cic:/matita/basics/logic/proj1.def(2)"proj1/a … (Hbisim a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e1,a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a ?)) lapply Hsub @(a href="cic:/matita/basics/list/list_elim_left.def(10)"list_elim_left/a … w) [//] #a #w1 #Hind #Hsub >a href="cic:/matita/tutorial/chapter9/moves_left.def(9)"moves_left/a >a href="cic:/matita/tutorial/chapter9/moves_left.def(9)"moves_left/a @(a href="cic:/matita/basics/logic/proj2.def(2)"proj2/a …(Hbisim …(Hind ?))) [#x #Hx @Hsub @a href="cic:/matita/tutorial/chapter5/memb_append_l1.def(5)"memb_append_l1/a // @@ -111,17 +111,17 @@ so we have just to check cofinality for any node we add to visited. Here is the extremely simple algorithm: *) -let rec bisim S l n (frontier,visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?) on n ≝ +img class="anchor" src="icons/tick.png" id="bisim"let rec bisim S l n (frontier,visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?) on n ≝ match n with - [ O ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visited〉 (* assert false *) + [ O ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visiteda title="Pair construction" href="cic:/fakeuri.def(1)"〉/a (* assert false *) | S m ⇒ match frontier with - [ nil ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a,visited〉 + [ nil ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a,visiteda title="Pair construction" href="cic:/fakeuri.def(1)"〉/a | cons hd tl ⇒ if a href="cic:/matita/tutorial/chapter4/beqb.def(2)"beqb/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a hd)) (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a hd)) then - bisim S l m (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? (λx.a href="cic:/matita/basics/bool/notb.def(1)"notb/a (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? x (hda title="cons" href="cic:/fakeuri.def(1)":/a:visited))) - (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a S l hd)) tl) (hda title="cons" href="cic:/fakeuri.def(1)":/a:visited) - else a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visited〉 + bisim S l m (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? (λx.a href="cic:/matita/basics/bool/notb.def(1)"notb/a (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? x (hda title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/avisited))) + (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a S l hd)) tl) (hda title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/avisited) + else a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visiteda title="Pair construction" href="cic:/fakeuri.def(1)"〉/a ] ]. @@ -133,43 +133,43 @@ if and only if all visited nodes are cofinal. The following results explicitly state the behaviour of bisim is the general case and in some relevant instances *) -lemma unfold_bisim: ∀S,l,n.∀frontier,visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. +img class="anchor" src="icons/tick.png" id="unfold_bisim"lemma unfold_bisim: ∀S,l,n.∀frontier,visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a S l n frontier visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a match n with - [ O ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visited〉 (* assert false *) + [ O ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visiteda title="Pair construction" href="cic:/fakeuri.def(1)"〉/a (* assert false *) | S m ⇒ match frontier with - [ nil ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a,visited〉 + [ nil ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a,visiteda title="Pair construction" href="cic:/fakeuri.def(1)"〉/a | cons hd tl ⇒ if a href="cic:/matita/tutorial/chapter4/beqb.def(2)"beqb/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a hd)) (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a hd)) then - a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a S l m (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? (λx.a href="cic:/matita/basics/bool/notb.def(1)"notb/a(a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? x (hda title="cons" href="cic:/fakeuri.def(1)":/a:visited))) - (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a S l hd)) tl) (hda title="cons" href="cic:/fakeuri.def(1)":/a:visited) - else a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visited〉 + a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a S l m (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? (λx.a href="cic:/matita/basics/bool/notb.def(1)"notb/a(a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? x (hda title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/avisited))) + (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a S l hd)) tl) (hda title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/avisited) + else a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visiteda title="Pair construction" href="cic:/fakeuri.def(1)"〉/a ] ]. #S #l #n cases n // qed. -lemma bisim_never: ∀S,l.∀frontier,visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. - a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a S l a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a frontier visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visited〉. +img class="anchor" src="icons/tick.png" id="bisim_never"lemma bisim_never: ∀S,l.∀frontier,visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. + a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a S l a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"O/a frontier visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visiteda title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. #frontier #visited >a href="cic:/matita/tutorial/chapter10/unfold_bisim.def(9)"unfold_bisim/a // qed. -lemma bisim_end: ∀Sig,l,m.∀visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. - a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a Sig l (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m) a title="nil" href="cic:/fakeuri.def(1)"[/a] visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a,visited〉. +img class="anchor" src="icons/tick.png" id="bisim_end"lemma bisim_end: ∀Sig,l,m.∀visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. + a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a Sig l (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m) a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a,visiteda title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. #n #visisted >a href="cic:/matita/tutorial/chapter10/unfold_bisim.def(9)"unfold_bisim/a // qed. -lemma bisim_step_true: ∀Sig,l,m.∀p.∀frontier,visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. +img class="anchor" src="icons/tick.png" id="bisim_step_true"lemma bisim_step_true: ∀Sig,l,m.∀p.∀frontier,visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. a href="cic:/matita/tutorial/chapter4/beqb.def(2)"beqb/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p)) (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → - a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a Sig l (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m) (pa title="cons" href="cic:/fakeuri.def(1)":/a:frontier) visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a - a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a Sig l m (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? (λx.a href="cic:/matita/basics/bool/notb.def(1)"notb/a(a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? x (pa title="cons" href="cic:/fakeuri.def(1)":/a:visited))) - (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a Sig l p)) frontier) (pa title="cons" href="cic:/fakeuri.def(1)":/a:visited). + a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a Sig l (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m) (pa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/afrontier) visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a + a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a Sig l m (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? (λx.a href="cic:/matita/basics/bool/notb.def(1)"notb/a(a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? x (pa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/avisited))) + (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a Sig l p)) frontier) (pa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/avisited). #Sig #l #m #p #frontier #visited #test >a href="cic:/matita/tutorial/chapter10/unfold_bisim.def(9)"unfold_bisim/a whd in ⊢ (??%?); >test // qed. -lemma bisim_step_false: ∀Sig,l,m.∀p.∀frontier,visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. +img class="anchor" src="icons/tick.png" id="bisim_step_false"lemma bisim_step_false: ∀Sig,l,m.∀p.∀frontier,visited: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. a href="cic:/matita/tutorial/chapter4/beqb.def(2)"beqb/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p)) (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → - a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a Sig l (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m) (pa title="cons" href="cic:/fakeuri.def(1)":/a:frontier) visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visited〉. + a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a Sig l (a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a m) (pa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/afrontier) visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,visiteda title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. #Sig #l #m #p #frontier #visited #test >a href="cic:/matita/tutorial/chapter10/unfold_bisim.def(9)"unfold_bisim/a whd in ⊢ (??%?); >test // qed. @@ -180,18 +180,18 @@ enumerate all possible pres: *) #b cases b normalize // qed. *) -let rec pitem_enum S (i:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S) on i ≝ +img class="anchor" src="icons/tick.png" id="pitem_enum"let rec pitem_enum S (i:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S) on i ≝ match i with - [ z ⇒ (a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"pz/a S)a title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a] - | e ⇒ (a href="cic:/matita/tutorial/chapter7/pitem.con(0,2,1)"pe/a S)a title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a] - | s y ⇒ (a href="cic:/matita/tutorial/chapter7/pitem.con(0,3,1)"ps/a S y)a title="cons" href="cic:/fakeuri.def(1)":/a:(a href="cic:/matita/tutorial/chapter7/pitem.con(0,4,1)"pp/a S y)a title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a] + [ z ⇒ (a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"pz/aspan class="error" title="Parse error: SYMBOL ':' or RPAREN expected after [term] (in [term])"/spanspan class="error" title="Parse error: SYMBOL ':' or RPAREN expected after [term] (in [term])"/span S)a title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a + | e ⇒ (a href="cic:/matita/tutorial/chapter7/pitem.con(0,2,1)"pe/a S)a title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a + | s y ⇒ (a href="cic:/matita/tutorial/chapter7/pitem.con(0,3,1)"ps/a S y)a title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/a(a href="cic:/matita/tutorial/chapter7/pitem.con(0,4,1)"pp/a S y)a title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a | o i1 i2 ⇒ a href="cic:/matita/basics/list/compose.def(2)"compose/a ??? (a href="cic:/matita/tutorial/chapter7/pitem.con(0,6,1)"po/a S) (pitem_enum S i1) (pitem_enum S i2) | c i1 i2 ⇒ a href="cic:/matita/basics/list/compose.def(2)"compose/a ??? (a href="cic:/matita/tutorial/chapter7/pitem.con(0,5,1)"pc/a S) (pitem_enum S i1) (pitem_enum S i2) | k i ⇒ a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a ?? (a href="cic:/matita/tutorial/chapter7/pitem.con(0,7,1)"pk/a S) (pitem_enum S i) ]. -lemma pitem_enum_complete : ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a (a href="cic:/matita/tutorial/chapter7/DeqItem.def(6)"DeqItem/a S) i (a href="cic:/matita/tutorial/chapter10/pitem_enum.fix(0,1,3)"pitem_enum/a S (a title="forget" href="cic:/fakeuri.def(1)"|/ai|)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +img class="anchor" src="icons/tick.png" id="pitem_enum_complete"lemma pitem_enum_complete : ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a (a href="cic:/matita/tutorial/chapter7/DeqItem.def(6)"DeqItem/a S) i (a href="cic:/matita/tutorial/chapter10/pitem_enum.fix(0,1,3)"pitem_enum/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aia title="forget" href="cic:/fakeuri.def(1)"|/a)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S #i elim i [1,2:// |3,4:#c normalize >(\b (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a … c)) // @@ -200,29 +200,29 @@ lemma pitem_enum_complete : ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pit ] qed. -definition pre_enum ≝ λS.λi:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S. - a href="cic:/matita/basics/list/compose.def(2)"compose/a ??? (λi,b.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,b〉) ( a href="cic:/matita/tutorial/chapter10/pitem_enum.fix(0,1,3)"pitem_enum/a S i) (a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="cons" href="cic:/fakeuri.def(1)":/a:a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a]). +img class="anchor" src="icons/tick.png" id="pre_enum"definition pre_enum ≝ λS.λi:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S. + a href="cic:/matita/basics/list/compose.def(2)"compose/a ??? (λi,b.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) (a href="cic:/matita/tutorial/chapter10/pitem_enum.fix(0,1,3)"pitem_enum/a S i) (a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a). -lemma pre_enum_complete : ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? e (a href="cic:/matita/tutorial/chapter10/pre_enum.def(4)"pre_enum/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. -#S * #i #b @(a href="cic:/matita/tutorial/chapter5/memb_compose.def(6)"memb_compose/a (a href="cic:/matita/tutorial/chapter7/DeqItem.def(6)"DeqItem/a S) a href="cic:/matita/tutorial/chapter4/DeqBool.def(5)"DeqBool/a ? (λi,b.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,b〉)) +img class="anchor" src="icons/tick.png" id="pre_enum_complete"lemma pre_enum_complete : ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? e (a href="cic:/matita/tutorial/chapter10/pre_enum.def(4)"pre_enum/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a ea title="forget" href="cic:/fakeuri.def(1)"|/a)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +#S * #i #b @(a href="cic:/matita/tutorial/chapter5/memb_compose.def(6)"memb_compose/a (a href="cic:/matita/tutorial/chapter7/DeqItem.def(6)"DeqItem/a S) a href="cic:/matita/tutorial/chapter4/DeqBool.def(5)"DeqBool/a ? (λi,b.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a)) // cases b normalize // qed. -definition space_enum ≝ λS.λi1,i2: a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S. - a href="cic:/matita/basics/list/compose.def(2)"compose/a ??? (λe1,e2.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ae1,e2〉) ( a href="cic:/matita/tutorial/chapter10/pre_enum.def(4)"pre_enum/a S i1) (a href="cic:/matita/tutorial/chapter10/pre_enum.def(4)"pre_enum/a S i2). +img class="anchor" src="icons/tick.png" id="space_enum"definition space_enum ≝ λS.λi1,i2:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S. + a href="cic:/matita/basics/list/compose.def(2)"compose/a ??? (λe1,e2.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ae1,e2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) (a href="cic:/matita/tutorial/chapter10/pre_enum.def(4)"pre_enum/a S i1) (a href="cic:/matita/tutorial/chapter10/pre_enum.def(4)"pre_enum/a S i2). -lemma space_enum_complete : ∀S.∀e1,e2: a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ae1,e2〉 ( a href="cic:/matita/tutorial/chapter10/space_enum.def(5)"space_enum/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1|) (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2|)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. -#S #e1 #e2 @(a href="cic:/matita/tutorial/chapter5/memb_compose.def(6)"memb_compose/a … (λi,b.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,b〉)) +img class="anchor" src="icons/tick.png" id="space_enum_complete"lemma space_enum_complete : ∀S.∀e1,e2: a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ae1,e2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a (a href="cic:/matita/tutorial/chapter10/space_enum.def(5)"space_enum/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1a title="forget" href="cic:/fakeuri.def(1)"|/a) (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2a title="forget" href="cic:/fakeuri.def(1)"|/a)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +#S #e1 #e2 @(a href="cic:/matita/tutorial/chapter5/memb_compose.def(6)"memb_compose/a … (λi,b.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a)) // qed. -definition all_reachable ≝ λS.λe1,e2: a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.λl: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. +img class="anchor" src="icons/tick.png" id="all_reachable"definition all_reachable ≝ λS.λe1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.λl: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. a href="cic:/matita/tutorial/chapter5/uniqueb.fix(0,1,5)"uniqueb/a ? l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="logical and" href="cic:/fakeuri.def(1)"∧/a ∀p. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? p l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a title="exists" href="cic:/fakeuri.def(1)"∃/aw.(a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p) a title="logical and" href="cic:/fakeuri.def(1)"∧/a (a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p). -definition disjoint ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λl1,l2. +img class="anchor" src="icons/tick.png" id="disjoint"definition disjoint ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λl1,l2. ∀p:S. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S p l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S p l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. (* We are ready to prove that bisim is correct; we use the invariant @@ -230,17 +230,17 @@ that at each call of bisim the two lists visited and frontier only contain nodes reachable from 〈e_1,e_2〉, hence it is absurd to suppose to meet a pair which is not cofinal. *) -lemma bisim_correct: ∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1}a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2} → +img class="anchor" src="icons/tick.png" id="bisim_correct"lemma bisim_correct: ∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1a title="in_prl" href="cic:/fakeuri.def(1)"}/aa title="extensional equality" href="cic:/fakeuri.def(1)"=/a1a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2a title="in_prl" href="cic:/fakeuri.def(1)"}/a → ∀l,n.∀frontier,visited:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ((a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)a title="Product" href="cic:/fakeuri.def(1)"×/a(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S)). - a title="norm" href="cic:/fakeuri.def(1)"|/aa href="cic:/matita/tutorial/chapter10/space_enum.def(5)"space_enum/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1|) (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2|)| a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n a title="natural plus" href="cic:/fakeuri.def(1)"+/a a title="norm" href="cic:/fakeuri.def(1)"|/avisited|→ + a title="norm" href="cic:/fakeuri.def(1)"|/aa href="cic:/matita/tutorial/chapter10/space_enum.def(5)"space_enum/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1a title="forget" href="cic:/fakeuri.def(1)"|/a) (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2a title="forget" href="cic:/fakeuri.def(1)"|/a)a title="norm" href="cic:/fakeuri.def(1)"|/a a title="natural 'less than'" href="cic:/fakeuri.def(1)"</a n a title="natural plus" href="cic:/fakeuri.def(1)"+/a a title="norm" href="cic:/fakeuri.def(1)"|/avisiteda title="norm" href="cic:/fakeuri.def(1)"|/a→ a href="cic:/matita/tutorial/chapter10/all_reachable.def(8)"all_reachable/a S e1 e2 visited → a href="cic:/matita/tutorial/chapter10/all_reachable.def(8)"all_reachable/a S e1 e2 frontier → a href="cic:/matita/tutorial/chapter10/disjoint.def(5)"disjoint/a ? frontier visited → a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a S l n frontier visited) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #Sig #e1 #e2 #same #l #n elim n [#frontier #visited #abs * #unique #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a @(a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a … abs) - @a href="cic:/matita/arithmetics/nat/le_to_not_lt.def(8)"le_to_not_lt/a @a href="cic:/matita/tutorial/chapter5/sublist_length.def(9)"sublist_length/a // * #e11 #e21 #membp - cut ((a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e11| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1|) a title="logical and" href="cic:/fakeuri.def(1)"∧/a (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e21| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2|)) + @a href="cic:/matita/arithmetics/nat/le_to_not_lt.def(8)"le_to_not_lt/a @a href="cic:/matita/tutorial/chapter5/sublist_length.def(6)"sublist_length/a // * #e11 #e21 #membp + cut ((|a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e11a title="forget" href="cic:/fakeuri.def(1)"|/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a |a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1a title="forget" href="cic:/fakeuri.def(1)"|/a) a title="logical and" href="cic:/fakeuri.def(1)"∧/a (|a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e21a title="forget" href="cic:/fakeuri.def(1)"|/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a |a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2a title="forget" href="cic:/fakeuri.def(1)"|/a)) [|* #H1 #H2

a href="cic:/matita/tutorial/chapter9/same_kernel_moves.def(9)"same_kernel_moves/a // |#m #HI * [#visited #vinv #finv >a href="cic:/matita/tutorial/chapter10/bisim_end.def(10)"bisim_end/a //] @@ -259,13 +259,13 @@ lemma bisim_correct: ∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.de |whd % [@a href="cic:/matita/tutorial/chapter5/unique_append_unique.def(6)"unique_append_unique/a @(a href="cic:/matita/basics/bool/andb_true_r.def(4)"andb_true_r/a … u_frontier)] @a href="cic:/matita/tutorial/chapter5/unique_append_elim.def(7)"unique_append_elim/a #q #H [cases (a href="cic:/matita/tutorial/chapter10/memb_sons.def(8)"memb_sons/a … (a href="cic:/matita/tutorial/chapter5/memb_filter_memb.def(5)"memb_filter_memb/a … H)) -H - #a * #m1 #m2 cases rp #w1 * #mw1 #mw2 @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … (w1a title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a]))) + #a * #m1 #m2 cases rp #w1 * #mw1 #mw2 @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … (w1a title="append" href="cic:/fakeuri.def(1)"@/a(a:a title="cons" href="cic:/fakeuri.def(1)":/a[a title="nil" href="cic:/fakeuri.def(1)"]/a))) >a href="cic:/matita/tutorial/chapter9/moves_left.def(9)"moves_left/a >a href="cic:/matita/tutorial/chapter9/moves_left.def(9)"moves_left/a >mw1 >mw2 >m1 >m2 % // |@r_frontier @a href="cic:/matita/tutorial/chapter5/memb_cons.def(5)"memb_cons/a // ] |@a href="cic:/matita/tutorial/chapter5/unique_append_elim.def(7)"unique_append_elim/a #q #H [@a href="cic:/matita/basics/bool/injective_notb.def(4)"injective_notb/a @(a href="cic:/matita/tutorial/chapter5/memb_filter_true.def(5)"memb_filter_true/a … H) - |cut ((qa title="eqb" href="cic:/fakeuri.def(1)"=/a=p) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a) + |cut ((q=a title="eqb" href="cic:/fakeuri.def(1)"=/ap) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a) [|#Hpq whd in ⊢ (??%?); >Hpq @disjoint @a href="cic:/matita/tutorial/chapter5/memb_cons.def(5)"memb_cons/a //] cases (a href="cic:/matita/basics/bool/andb_true.def(5)"andb_true/a … u_frontier) #notp #_ @(\bf ?) @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … a href="cic:/matita/basics/bool/not_eq_true_false.def(3)"not_eq_true_false/a) #eqqp H // @@ -278,17 +278,17 @@ qed. and the sons of visited are either in visited or in the frontier; since at the end frontier is empty, visited is hence a bisimulation. *) -definition all_true ≝ λS.λl.∀p:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S) a title="Product" href="cic:/fakeuri.def(1)"×/a (a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? p l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → +img class="anchor" src="icons/tick.png" id="all_true"definition all_true ≝ λS.λl.∀p:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S) a title="Product" href="cic:/fakeuri.def(1)"×/a (a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? p l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → (a href="cic:/matita/tutorial/chapter4/beqb.def(2)"beqb/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p)) (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a). -definition sub_sons ≝ λS,l,l1,l2.∀x:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S) a title="Product" href="cic:/fakeuri.def(1)"×/a (a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). +img class="anchor" src="icons/tick.png" id="sub_sons"definition sub_sons ≝ λS,l,l1,l2.∀x:(a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S) a title="Product" href="cic:/fakeuri.def(1)"×/a (a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S). a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? x l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a ? (a href="cic:/matita/tutorial/chapter10/sons.def(7)"sons/a ? l x) l2. -lemma bisim_complete: +img class="anchor" src="icons/tick.png" id="bisim_complete"lemma bisim_complete: ∀S,l,n.∀frontier,visited,visited_res:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a ?. a href="cic:/matita/tutorial/chapter10/all_true.def(8)"all_true/a S visited → a href="cic:/matita/tutorial/chapter10/sub_sons.def(8)"sub_sons/a S l visited (frontiera title="append" href="cic:/fakeuri.def(1)"@/avisited) → - a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a S l n frontier visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a,visited_res〉 → + a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a S l n frontier visited a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a 〈a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a,visited_resa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a → a href="cic:/matita/tutorial/chapter10/is_bisim.def(8)"is_bisim/a S visited_res l a title="logical and" href="cic:/fakeuri.def(1)"∧/a a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a ? (frontiera title="append" href="cic:/fakeuri.def(1)"@/avisited) visited_res. #S #l #n elim n [#fron #vis #vis_res #_ #_ >a href="cic:/matita/tutorial/chapter10/bisim_never.def(10)"bisim_never/a #H destruct @@ -303,7 +303,7 @@ lemma bisim_complete: (* frontier = hd:: tl and hd is ok *) #H #tl #visited #visited_res #allv >(a href="cic:/matita/tutorial/chapter10/bisim_step_true.def(10)"bisim_step_true/a … H) (* new_visited = hd::visited are all ok *) - cut (a href="cic:/matita/tutorial/chapter10/all_true.def(8)"all_true/a S (hda title="cons" href="cic:/fakeuri.def(1)":/a:visited)) + cut (a href="cic:/matita/tutorial/chapter10/all_true.def(8)"all_true/a S (hd:a title="cons" href="cic:/fakeuri.def(1)":/avisited)) [#p #H1 cases (a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a … H1) [#eqp >(\P eqp) @H |@allv]] (* we now exploit the induction hypothesis *) #allh #subH #bisim cases (Hind … allh … bisim) -bisim -Hind @@ -320,7 +320,7 @@ lemma bisim_complete: #eqhdx <(\P eqhdx) #xa #membxa (* xa is a son of x; we must distinguish the case xa was already visited form the case xa is new *) - cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? xa (xa title="cons" href="cic:/fakeuri.def(1)":/a:visited))) + cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a ? xa (x:a title="cons" href="cic:/fakeuri.def(1)":/avisited))) [(* xa visited - trivial *) #membxa @a href="cic:/matita/tutorial/chapter5/memb_append_l2.def(5)"memb_append_l2/a // |(* xa new *) #membxa @a href="cic:/matita/tutorial/chapter5/memb_append_l1.def(5)"memb_append_l1/a @a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l1.def(6)"sublist_unique_append_l1/a @a href="cic:/matita/tutorial/chapter5/memb_filter_l.def(5)"memb_filter_l/a [>membxa //|//] @@ -342,18 +342,18 @@ qed. prove that two expressions are equivalente if and only if they define the same language. *) -definition equiv ≝ λSig.λre1,re2:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a Sig. +img class="anchor" src="icons/tick.png" id="equiv"definition equiv ≝ λSig.λre1,re2:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a Sig. let e1 ≝ a title="eclose" href="cic:/fakeuri.def(1)"•/a(a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a ? re1) in let e2 ≝ a title="eclose" href="cic:/fakeuri.def(1)"•/a(a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a ? re2) in - let n ≝ a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a (a href="cic:/matita/basics/list/length.fix(0,1,1)"length/a ? (a href="cic:/matita/tutorial/chapter10/space_enum.def(5)"space_enum/a Sig (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1|) (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2|))) in + let n ≝ a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"S/a (a href="cic:/matita/basics/list/length.fix(0,1,1)"length/a ? (a href="cic:/matita/tutorial/chapter10/space_enum.def(5)"space_enum/a Sig (|a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1a title="forget" href="cic:/fakeuri.def(1)"|/a) (|a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2a title="forget" href="cic:/fakeuri.def(1)"|/a))) in let sig ≝ (a href="cic:/matita/tutorial/chapter10/occ.def(7)"occ/a Sig e1 e2) in - (a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a ? sig n (a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ae1,e2〉a title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a]) a title="nil" href="cic:/fakeuri.def(1)"[/a]). + (a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"bisim/a ? sig n (〈e1,e2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a:a title="cons" href="cic:/fakeuri.def(1)":/a[a title="nil" href="cic:/fakeuri.def(1)"]/a) [a title="nil" href="cic:/fakeuri.def(1)"]/a). -theorem euqiv_sem : ∀Sig.∀e1,e2:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a Sig. - a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter10/equiv.def(9)"equiv/a ? e1 e2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="iff" href="cic:/fakeuri.def(1)"↔/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{e1} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{e2}. +img class="anchor" src="icons/tick.png" id="euqiv_sem"theorem euqiv_sem : ∀Sig.∀e1,e2:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a Sig. + a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter10/equiv.def(9)"equiv/a ? e1 e2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="iff" href="cic:/fakeuri.def(1)"↔/a \sem{e1a title="in_l" href="cic:/fakeuri.def(1)"}/a =1 \sem{e2a title="in_l" href="cic:/fakeuri.def(1)"}/a. #Sig #re1 #re2 % [#H @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter8/re_embedding.def(13)"re_embedding/a] @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter8/re_embedding.def(13)"re_embedding/a] - cut (a href="cic:/matita/tutorial/chapter10/equiv.def(9)"equiv/a ? re1 re2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a,a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a href="cic:/matita/tutorial/chapter10/equiv.def(9)"equiv/a ? re1 re2)〉) + cut (a href="cic:/matita/tutorial/chapter10/equiv.def(9)"equiv/a ? re1 re2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a 〈a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a,a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a href="cic:/matita/tutorial/chapter10/equiv.def(9)"equiv/a ? re1 re2)a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) [H normalize // qed. -lemma div2S1: ∀n,q. a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aq,a href="cic:/matita/tutorial/chapter2/bool.con(0,1,0)"tt/a〉 → a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a n) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aspan class="error" title="Parse error: [term] expected after [sym=] (in [term])"/span a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a q,a href="cic:/matita/tutorial/chapter2/bool.con(0,2,0)"ff/a〉. +img class="anchor" src="icons/tick.png" id="div2S1"lemma div2S1: ∀n,q. a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aq,a href="cic:/matita/tutorial/chapter2/bool.con(0,1,0)"tt/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a → a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a n) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aspan class="error" title="Parse error: [term] expected after [sym=] (in [term])"/span a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a q,a href="cic:/matita/tutorial/chapter2/bool.con(0,2,0)"ff/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. #n #q #H normalize >H normalize // qed. -lemma div2_ok: ∀n,q,r. a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aq,r〉 → n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter2/add.fix(0,0,1)"add/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a q) (a href="cic:/matita/tutorial/chapter2/nat_of_bool.def(1)"nat_of_bool/a r). +img class="anchor" src="icons/tick.png" id="div2_ok"lemma div2_ok: ∀n,q,r. a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aq,ra title="Pair construction" href="cic:/fakeuri.def(1)"〉/a → n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter2/add.fix(0,0,1)"add/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a q) (a href="cic:/matita/tutorial/chapter2/nat_of_bool.def(1)"nat_of_bool/a r). #n elim n [#q #r normalize #H destruct // |#a #Hind #q #r - cut (a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/types/fst.fix(0,2,1)"fst/a … (a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a a), a href="cic:/matita/basics/types/snd.fix(0,2,1)"snd/a … (a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a a)〉) [//] + cut (a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/types/fst.fix(0,2,1)"fst/a … (a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a a), a href="cic:/matita/basics/types/snd.fix(0,2,1)"snd/a … (a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a a)a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) [//] cases (a href="cic:/matita/basics/types/snd.fix(0,2,1)"snd/a … (a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a a)) [#H >(a href="cic:/matita/tutorial/chapter2/div2S1.def(3)"div2S1/a … H) #H1 destruct @a href="cic:/matita/basics/logic/eq_f.def(3)"eq_f/a span style="text-decoration: underline;">/spana href="cic:/matita/tutorial/chapter2/add_S.def(2)"add_S/a whd in ⊢ (???%); <a href="cic:/matita/tutorial/chapter2/add_S.def(2)"add_S/a @(Hind … H) |#H >(a href="cic:/matita/tutorial/chapter2/div2SO.def(3)"div2SO/a … H) #H1 destruct >a href="cic:/matita/tutorial/chapter2/add_S.def(2)"add_S/a @a href="cic:/matita/basics/logic/eq_f.def(3)"eq_f/a @(Hind … H) @@ -237,7 +237,7 @@ qed. h2 class="section"Mixing proofs and computations/h2 There is still another possibility, however, namely to mix the program and its specification into a single entity. The idea is to refine the output type of the -div2 function: it should not be just a generic pair 〈q,r〉 of natural numbers but a +div2 function: it should not be just a generic pair 〈q,r〉 of natural numbers but a specific pair satisfying the specification of the function. In other words, we need the possibility to define, for a type A and a property P over A, the subset type {a:A|P(a)} of all elements a of type A that satisfy the property P. Subset types @@ -252,36 +252,36 @@ by an element of a of type A and a proof of P(a). The crucial point is to have a language reach enough to comprise proofs among its expressions. *) -record Sub (A:Type[0]) (P:A → Prop) : Type[0] ≝ +img class="anchor" src="icons/tick.png" id="Sub"record Sub (A:Type[0]) (P:A → Prop) : Type[0] ≝ {witness: A; proof: P witness}. -definition qr_spec ≝ λn.λp.∀q,r. p a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aq,r〉 → n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter2/add.fix(0,0,1)"add/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a q) (a href="cic:/matita/tutorial/chapter2/nat_of_bool.def(1)"nat_of_bool/a r). +img class="anchor" src="icons/tick.png" id="qr_spec"definition qr_spec ≝ λn.λp.∀q,r. p a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aq,ra title="Pair construction" href="cic:/fakeuri.def(1)"〉/a → n a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter2/add.fix(0,0,1)"add/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a q) (a href="cic:/matita/tutorial/chapter2/nat_of_bool.def(1)"nat_of_bool/a r). (* We can now construct a function from n to {p|qr_spec n p} by composing the objects we already have *) -definition div2P: ∀n. a href="cic:/matita/tutorial/chapter2/Sub.ind(1,0,2)"Sub/a (a href="cic:/matita/tutorial/chapter2/nat.ind(1,0,0)"nat/aa title="Product" href="cic:/fakeuri.def(1)"×/aspan style="text-decoration: underline;"a href="cic:/matita/tutorial/chapter2/bool.ind(1,0,0)"bool/a/span) (a href="cic:/matita/tutorial/chapter2/qr_spec.def(3)"qr_spec/a n) ≝ λn. +img class="anchor" src="icons/tick.png" id="div2P"definition div2P: ∀n. a href="cic:/matita/tutorial/chapter2/Sub.ind(1,0,2)"Sub/a (a href="cic:/matita/tutorial/chapter2/nat.ind(1,0,0)"nat/aa title="Product" href="cic:/fakeuri.def(1)"×/aspan style="text-decoration: underline;"a href="cic:/matita/tutorial/chapter2/bool.ind(1,0,0)"bool/a/span) (a href="cic:/matita/tutorial/chapter2/qr_spec.def(3)"qr_spec/a n) ≝ λn. a href="cic:/matita/tutorial/chapter2/Sub.con(0,1,2)"mk_Sub/a ?? (a href="cic:/matita/tutorial/chapter2/div2.fix(0,0,2)"div2/a n) (a href="cic:/matita/tutorial/chapter2/div2_ok.def(4)"div2_ok/a n). (* But we can also try do directly build such an object *) -definition div2Pagain : ∀n.a href="cic:/matita/tutorial/chapter2/Sub.ind(1,0,2)"Sub/a (a href="cic:/matita/tutorial/chapter2/nat.ind(1,0,0)"nat/aa title="Product" href="cic:/fakeuri.def(1)"×/aspan style="text-decoration: underline;"/spana href="cic:/matita/tutorial/chapter2/bool.ind(1,0,0)"bool/a) (a href="cic:/matita/tutorial/chapter2/qr_spec.def(3)"qr_spec/a n). +img class="anchor" src="icons/tick.png" id="div2Pagain"definition div2Pagain : ∀n.a href="cic:/matita/tutorial/chapter2/Sub.ind(1,0,2)"Sub/a (a href="cic:/matita/tutorial/chapter2/nat.ind(1,0,0)"nat/aa title="Product" href="cic:/fakeuri.def(1)"×/aspan style="text-decoration: underline;"/spana href="cic:/matita/tutorial/chapter2/bool.ind(1,0,0)"bool/a) (a href="cic:/matita/tutorial/chapter2/qr_spec.def(3)"qr_spec/a n). #n elim n - [@(a href="cic:/matita/tutorial/chapter2/Sub.con(0,1,2)"mk_Sub/a … a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a,a href="cic:/matita/tutorial/chapter2/bool.con(0,2,0)"ff/a〉) normalize #q #r #H destruct // + [@(a href="cic:/matita/tutorial/chapter2/Sub.con(0,1,2)"mk_Sub/a … a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a,a href="cic:/matita/tutorial/chapter2/bool.con(0,2,0)"ff/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) normalize #q #r #H destruct // |#a * #p #qrspec - cut (p a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/types/fst.fix(0,2,1)"fst/a … p, a href="cic:/matita/basics/types/snd.fix(0,2,1)"snd/a … p〉) [//] + cut (p a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/types/fst.fix(0,2,1)"fst/a … p, a href="cic:/matita/basics/types/snd.fix(0,2,1)"snd/a … pa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) [//] cases (a href="cic:/matita/basics/types/snd.fix(0,2,1)"snd/a … p) - [#H @(a href="cic:/matita/tutorial/chapter2/Sub.con(0,1,2)"mk_Sub/a … a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a (a href="cic:/matita/basics/types/fst.fix(0,2,1)"fst/a … p),a href="cic:/matita/tutorial/chapter2/bool.con(0,2,0)"ff/a〉) whd #q #r #H1 destruct @a href="cic:/matita/basics/logic/eq_f.def(3)"eq_f/a span style="text-decoration: underline;">/spana href="cic:/matita/tutorial/chapter2/add_S.def(2)"add_S/a + [#H @(a href="cic:/matita/tutorial/chapter2/Sub.con(0,1,2)"mk_Sub/a … a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a (a href="cic:/matita/basics/types/fst.fix(0,2,1)"fst/a … p),a href="cic:/matita/tutorial/chapter2/bool.con(0,2,0)"ff/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) whd #q #r #H1 destruct @a href="cic:/matita/basics/logic/eq_f.def(3)"eq_f/a span style="text-decoration: underline;">/spana href="cic:/matita/tutorial/chapter2/add_S.def(2)"add_S/a whd in ⊢ (???%); <a href="cic:/matita/tutorial/chapter2/add_S.def(2)"add_S/a @(qrspec … H) - |#H @(a href="cic:/matita/tutorial/chapter2/Sub.con(0,1,2)"mk_Sub/a … a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/types/fst.fix(0,2,1)"fst/a … p,a href="cic:/matita/tutorial/chapter2/bool.con(0,1,0)"tt/a〉) whd #q #r #H1 destruct >a href="cic:/matita/tutorial/chapter2/add_S.def(2)"add_S/a @a href="cic:/matita/basics/logic/eq_f.def(3)"eq_f/a @(qrspec … H) + |#H @(a href="cic:/matita/tutorial/chapter2/Sub.con(0,1,2)"mk_Sub/a … a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/types/fst.fix(0,2,1)"fst/a … p,a href="cic:/matita/tutorial/chapter2/bool.con(0,1,0)"tt/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) whd #q #r #H1 destruct >a href="cic:/matita/tutorial/chapter2/add_S.def(2)"add_S/a @a href="cic:/matita/basics/logic/eq_f.def(3)"eq_f/a @(qrspec … H) ] qed. -example quotient7: a href="cic:/matita/tutorial/chapter2/witness.fix(0,2,1)"witness/a … (a href="cic:/matita/tutorial/chapter2/div2Pagain.def(4)"div2Pagain/a (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a)))))))) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a)),a href="cic:/matita/tutorial/chapter2/bool.con(0,1,0)"tt/a〉. +img class="anchor" src="icons/tick.png" id="quotient7"example quotient7: a href="cic:/matita/tutorial/chapter2/witness.fix(0,2,1)"witness/a … (a href="cic:/matita/tutorial/chapter2/div2Pagain.def(4)"div2Pagain/a (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a)))))))) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a(a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a)),a href="cic:/matita/tutorial/chapter2/bool.con(0,1,0)"tt/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. // qed. -example quotient8: a href="cic:/matita/tutorial/chapter2/witness.fix(0,2,1)"witness/a … (a href="cic:/matita/tutorial/chapter2/div2Pagain.def(4)"div2Pagain/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a)))))) - a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a))), a href="cic:/matita/tutorial/chapter2/bool.con(0,2,0)"ff/a〉. +img class="anchor" src="icons/tick.png" id="quotient8"example quotient8: a href="cic:/matita/tutorial/chapter2/witness.fix(0,2,1)"witness/a … (a href="cic:/matita/tutorial/chapter2/div2Pagain.def(4)"div2Pagain/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a)))))) + a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/twice.def(2)"twice/a (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a))), a href="cic:/matita/tutorial/chapter2/bool.con(0,2,0)"ff/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. // qed. -prepre /pre/pre +prepre /pre/pre \ No newline at end of file diff --git a/weblib/tutorial/chapter3.ma b/weblib/tutorial/chapter3.ma index 51b810bee..c6cb5b8cd 100644 --- a/weblib/tutorial/chapter3.ma +++ b/weblib/tutorial/chapter3.ma @@ -7,7 +7,7 @@ include "basics/bool.ma". (* Matita supports polymorphic data types. The most typical case are polymorphic lists, parametric in the type of their elements: *) -inductive list (A:Type[0]) : Type[0] ≝ +img class="anchor" src="icons/tick.png" id="list"inductive list (A:Type[0]) : Type[0] ≝ | nil: list A | cons: A -> list A -> list A. @@ -64,17 +64,17 @@ specify in the its defintion on which one of them we are recurring: in this case If not othewise specified, recursion is supposed to act on the first argument of the function.*) -let rec append A (l1: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) l2 on l1 ≝ +img class="anchor" src="icons/tick.png" id="append"let rec append A (l1: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) l2 on l1 ≝ match l1 with [ nil ⇒ l2 - | cons hd tl ⇒ hd a title="cons" href="cic:/fakeuri.def(1)":/aspan class="error" title="Parse error: [sym:] expected after [sym:] (in [term])"/span: append A tl l2 ]. + | cons hd tl ⇒ hd a title="cons" href="cic:/fakeuri.def(1)":/aspan class="error" title="Parse error: [sym:] expected after [sym:] (in [term])"/spana title="cons" href="cic:/fakeuri.def(1)":/a append A tl l2 ]. interpretation "append" 'append l1 l2 = (append ? l1 l2). (* As usual, the function is executable. For instance, (append A nil l) reduces to l, as shown by the following example: *) -example nil_append: ∀A.∀l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A. a title="nil" href="cic:/fakeuri.def(1)"[/aspan class="error" title="Parse error: [term] expected after [sym[] (in [term])"/span] a title="append" href="cic:/fakeuri.def(1)"@/a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a l. +img class="anchor" src="icons/tick.png" id="nil_append"example nil_append: ∀A.∀l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A. a title="nil" href="cic:/fakeuri.def(1)"[/aspan class="error" title="Parse error: [term] expected after [sym[] (in [term])"/spana title="nil" href="cic:/fakeuri.def(1)"]/a a title="append" href="cic:/fakeuri.def(1)"@/a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a l. #A #l normalize // qed. (* Proving that l @ [] = l is just a bit more complex. The situation is exactly @@ -82,7 +82,7 @@ the same as for the addition operation of the previous chapter: since append is defined by recutsion over the first argument, the computation of l @ [] is stuck, and we must proceed by induction on l *) -lemma append_nil: ∀A.∀l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A.l a title="append" href="cic:/fakeuri.def(1)"@/aspan class="error" title="Parse error: [term level 46] expected after [sym@] (in [term])"/spanspan class="error" title="Parse error: [term level 46] expected after [sym@] (in [term])"/span a title="nil" href="cic:/fakeuri.def(1)"[/a] a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a l. +img class="anchor" src="icons/tick.png" id="append_nil"lemma append_nil: ∀A.∀l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A.l a title="append" href="cic:/fakeuri.def(1)"@/aspan class="error" title="Parse error: [term level 46] expected after [sym@] (in [term])"/spanspan class="error" title="Parse error: [term level 46] expected after [sym@] (in [term])"/span a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a l. #A #l (elim l) normalize // qed. (* similarly, we can define the two functions head and tail. Since we can only define @@ -90,53 +90,53 @@ total functions, we should decide what to do in case the input list is empty. For tl, it is natural to return the empty list; for hd, we take in input a default element d of type A to return in this case. *) -definition head ≝ λA.λl: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A.λd:A. +img class="anchor" src="icons/tick.png" id="head"definition head ≝ λA.λl: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A.λd:A. match l with [ nil ⇒ d | cons a _ ⇒ a]. -definition tail ≝ λA.λl: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A. - match l with [ nil ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/a] | cons hd tl ⇒ tl]. +img class="anchor" src="icons/tick.png" id="tail"definition tail ≝ λA.λl: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A. + match l with [ nil ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a | cons hd tl ⇒ tl]. -example ex_head: ∀A.∀a,d,l. a href="cic:/matita/tutorial/chapter3/head.def(1)"head/a A (aa title="cons" href="cic:/fakeuri.def(1)":/a:l) d a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aspan class="error" title="Parse error: [term] expected after [sym=] (in [term])"/spanspan class="error" title="Parse error: [term] expected after [sym=] (in [term])"/span a. +img class="anchor" src="icons/tick.png" id="ex_head"example ex_head: ∀A.∀a,d,l. a href="cic:/matita/tutorial/chapter3/head.def(1)"head/a A (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al) d a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aspan class="error" title="Parse error: [term] expected after [sym=] (in [term])"/spanspan class="error" title="Parse error: [term] expected after [sym=] (in [term])"/span a. #A #a #d #l normalize // qed. -example ex_tail: a href="cic:/matita/tutorial/chapter3/tail.def(1)"tail/a ? (a href="cic:/matita/tutorial/chapter3/list.con(0,2,1)"cons/a ? a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="nil" href="cic:/fakeuri.def(1)"[/a]) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/a]. +img class="anchor" src="icons/tick.png" id="ex_tail"example ex_tail: a href="cic:/matita/tutorial/chapter3/tail.def(1)"tail/a ? (a href="cic:/matita/tutorial/chapter3/list.con(0,2,1)"cons/a ? a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a. normalize // qed. -theorem associative_append: +img class="anchor" src="icons/tick.png" id="associative_append"theorem associative_append: ∀A.∀l1,l2,l3: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A. (l1 a title="append" href="cic:/fakeuri.def(1)"@/a l2) a title="append" href="cic:/fakeuri.def(1)"@/a l3 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a l1 a title="append" href="cic:/fakeuri.def(1)"@/a (l2 a title="append" href="cic:/fakeuri.def(1)"@/a l3). #A #l1 #l2 #l3 (elim l1) normalize // qed. (* Problemi con la notazione *) -lemma a_append: ∀A.∀a.∀l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A. (aa title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a]) a title="append" href="cic:/fakeuri.def(1)"@/a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a aa title="cons" href="cic:/fakeuri.def(1)":/a:l. +img class="anchor" src="icons/tick.png" id="a_append"lemma a_append: ∀A.∀a.∀l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A. (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a) a title="append" href="cic:/fakeuri.def(1)"@/a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al. // qed. -theorem append_cons: -∀A.∀a:A.∀l,l1: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A.la title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/a:l1)a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (l a title="append" href="cic:/fakeuri.def(1)"@/a (a href="cic:/matita/tutorial/chapter3/list.con(0,2,1)"cons/a ? a a title="nil" href="cic:/fakeuri.def(1)"[/a])) a title="append" href="cic:/fakeuri.def(1)"@/a l1. +img class="anchor" src="icons/tick.png" id="append_cons"theorem append_cons: +∀A.∀a:A.∀l,l1: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A.la title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al1)a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (l a title="append" href="cic:/fakeuri.def(1)"@/a (a href="cic:/matita/tutorial/chapter3/list.con(0,2,1)"cons/a ? a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a)) a title="append" href="cic:/fakeuri.def(1)"@/a l1. // qed. (* Other typical functions over lists are those computing the length of a list, and the function returning the nth element *) -let rec length (A:Type[0]) (l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) on l ≝ +img class="anchor" src="icons/tick.png" id="length"let rec length (A:Type[0]) (l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) on l ≝ match l with [ nil ⇒ a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a | cons a tl ⇒ a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a (length A tl)]. -let rec nth n (A:Type[0]) (l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) (d:A) ≝ +img class="anchor" src="icons/tick.png" id="nth"let rec nth n (A:Type[0]) (l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) (d:A) ≝ match n with [O ⇒ a href="cic:/matita/tutorial/chapter3/head.def(1)"head/a A l d |S m ⇒ nth m A (a href="cic:/matita/tutorial/chapter3/tail.def(1)"tail/a A l) d]. -example ex_length: a href="cic:/matita/tutorial/chapter3/length.fix(0,1,1)"length/a ? (a href="cic:/matita/tutorial/chapter3/list.con(0,2,1)"cons/a ? a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a a title="nil" href="cic:/fakeuri.def(1)"[/a]) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a. +img class="anchor" src="icons/tick.png" id="ex_length"example ex_length: a href="cic:/matita/tutorial/chapter3/length.fix(0,1,1)"length/a ? (a href="cic:/matita/tutorial/chapter3/list.con(0,2,1)"cons/a ? a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a. normalize // qed. -example ex_nth: a href="cic:/matita/tutorial/chapter3/nth.fix(0,0,2)"nth/a (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a) ? (a href="cic:/matita/tutorial/chapter3/list.con(0,2,1)"cons/a ? (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a) (a href="cic:/matita/tutorial/chapter3/list.con(0,2,1)"cons/a ? a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a a title="nil" href="cic:/fakeuri.def(1)"[/aspan class="error" title="Parse error: [term] expected after [sym[] (in [term])"/span])) a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a. +img class="anchor" src="icons/tick.png" id="ex_nth"example ex_nth: a href="cic:/matita/tutorial/chapter3/nth.fix(0,0,2)"nth/a (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a) ? (a href="cic:/matita/tutorial/chapter3/list.con(0,2,1)"cons/a ? (a href="cic:/matita/tutorial/chapter2/nat.con(0,2,0)"S/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a) (a href="cic:/matita/tutorial/chapter3/list.con(0,2,1)"cons/a ? a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a a title="nil" href="cic:/fakeuri.def(1)"[/aspan class="error" title="Parse error: [term] expected after [sym[] (in [term])"/spana title="nil" href="cic:/fakeuri.def(1)"]/a)) a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter2/nat.con(0,1,0)"O/a. normalize // qed. (* Proving that the length of l1@l2 is the sum of the lengths of l1 and l2 just requires a trivial induction on the first list. *) - lemma length_add: ∀A.∀l1,l2:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A. + img class="anchor" src="icons/tick.png" id="length_add"lemma length_add: ∀A.∀l1,l2:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A. a href="cic:/matita/tutorial/chapter3/length.fix(0,1,1)"length/a ? (l1a title="append" href="cic:/fakeuri.def(1)"@/al2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter2/add.fix(0,0,1)"add/a (a href="cic:/matita/tutorial/chapter3/length.fix(0,1,1)"length/a ? l1) (a href="cic:/matita/tutorial/chapter3/length.fix(0,1,1)"length/a ? l2). #A #l1 elim l1 normalize // qed. @@ -147,22 +147,22 @@ list is different from any list with at least one element, that is from any list of the kind (a::l)? We start defining a simple predicate stating if a list is empty or not. The predicate is computed by inspection over the list *) -definition is_nil: ∀A:Type[0].a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A → Prop ≝ -λA.λl.match l with [ nil ⇒ l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/a] | cons hd tl ⇒ (l a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a title="nil" href="cic:/fakeuri.def(1)"[/a])]. +img class="anchor" src="icons/tick.png" id="is_nil"definition is_nil: ∀A:Type[0].a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A → Prop ≝ +λA.λl.match l with [ nil ⇒ l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a | cons hd tl ⇒ (l a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a)]. (* Next we need a simple result about negation: if you wish to prove ¬P you are authorized to add P to your hypothesis: *) -lemma neg_aux : ∀P:Prop. (P → a title="logical not" href="cic:/fakeuri.def(1)"¬/aP) → a title="logical not" href="cic:/fakeuri.def(1)"¬/aP. +img class="anchor" src="icons/tick.png" id="neg_aux"lemma neg_aux : ∀P:Prop. (P → a title="logical not" href="cic:/fakeuri.def(1)"¬/aP) → a title="logical not" href="cic:/fakeuri.def(1)"¬/aP. #P #PtonegP % /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ qed. -theorem diff_cons_nil: -∀A:Type[0].∀l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A.∀a:A. aa title="cons" href="cic:/fakeuri.def(1)":/a:l a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a title="nil" href="cic:/fakeuri.def(1)"[/a]. +img class="anchor" src="icons/tick.png" id="diff_cons_nil"theorem diff_cons_nil: +∀A:Type[0].∀l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A.∀a:A. aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a. #A #l #a @a href="cic:/matita/tutorial/chapter3/neg_aux.def(3)"neg_aux/a #Heq (* we start assuming the new hypothesis Heq of type a::l = [] using neg_aux. Next we use the change tactic to pass from the current goal a::l≠ [] to the expression is_nil a::l, convertible with it. *) -(change with (a href="cic:/matita/tutorial/chapter3/is_nil.def(1)"is_nil/a ? (aa title="cons" href="cic:/fakeuri.def(1)":/a:l))) +(change with (a href="cic:/matita/tutorial/chapter3/is_nil.def(1)"is_nil/a ? (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al))) (* Now, we rewrite with Heq, obtaining (is_nil A []), that reduces to the trivial goal [] = [] *) >Heq // qed. @@ -177,8 +177,8 @@ False_ind: ∀P.False → P to the current goal, that breaks down to prove False then absurd: ∀A:Prop. A → ¬A → False to reduce to the contradictory cases. Usually, you may invoke automation to take care to solve the absurd case. *) -lemma nil_to_nil: ∀A.∀l1,l2:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a span style="text-decoration: underline;"/spanA. - l1a title="append" href="cic:/fakeuri.def(1)"@/al2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/a] → l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/a] a title="logical and" href="cic:/fakeuri.def(1)"∧/a l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/a]. +img class="anchor" src="icons/tick.png" id="nil_to_nil"lemma nil_to_nil: ∀A.∀l1,l2:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a span style="text-decoration: underline;"/spanA. + l1a title="append" href="cic:/fakeuri.def(1)"@/al2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a → l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a a title="logical and" href="cic:/fakeuri.def(1)"∧/a l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a. #A #l1 cases l1 normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ #a #tl #l2 #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ qed. (* @@ -188,14 +188,14 @@ acting over lists. A typical example is the map function, taking a function f:A → B, a list l = [a1; a2; ... ; an] and returning the list [f a1; f a2; ... ; f an]. *) -let rec map (A,B:Type[0]) (f: A → B) (l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) on l: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a B ≝ - match l with [ nil ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/a] | cons x tl ⇒ f x a title="cons" href="cic:/fakeuri.def(1)":/a: (map A B f tl)]. +img class="anchor" src="icons/tick.png" id="map"let rec map (A,B:Type[0]) (f: A → B) (l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) on l: a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a B ≝ + match l with [ nil ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a | cons x tl ⇒ f x a title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/a (map A B f tl)]. (* Another major example is the fold function, that taken a list l = [a1; a2; ... ;an], a base value b:B, and a function f: A → B → B returns (f a1 (f a2 (... (f an b)...))). *) -let rec foldr (A,B:Type[0]) (f:A → B → B) (b:B) (l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) on l :B ≝ +img class="anchor" src="icons/tick.png" id="foldr"let rec foldr (A,B:Type[0]) (f:A → B → B) (b:B) (l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) on l :B ≝ match l with [ nil ⇒ b | cons a l ⇒ f a (foldr A B f b l)]. (* As an example of application of foldr, let us use it to define a filter @@ -205,21 +205,21 @@ foldr should be (list A), the base value is [], and f: A → list A →list A is the function that taken x and l returns x::l, if x satisfies the test, and l otherwise. We use an if_then_else function included from bool.ma to this purpose. *) -definition filter ≝ +img class="anchor" src="icons/tick.png" id="filter"definition filter ≝ λT.λp:T → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. - a href="cic:/matita/tutorial/chapter3/foldr.fix(0,4,1)"foldr/a T (a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a T) (λx,l0. if p x then xa title="cons" href="cic:/fakeuri.def(1)":/a:l0 else l0) a title="nil" href="cic:/fakeuri.def(1)"[/a]. + a href="cic:/matita/tutorial/chapter3/foldr.fix(0,4,1)"foldr/a T (a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a T) (λx,l0. if p x then xa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al0 else l0) a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a. (* Here are a couple of simple lemmas on the behaviour of the filter function. It is often convenient to state such lemmas, in order to be able to use rewriting as an alternative to reduction in proofs: reduction is a bit difficult to control. *) -lemma filter_true : ∀A,l,a,p. p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → - a href="cic:/matita/tutorial/chapter3/filter.def(2)"filter/a A p (aa title="cons" href="cic:/fakeuri.def(1)":/a:l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a a title="cons" href="cic:/fakeuri.def(1)":/a: a href="cic:/matita/tutorial/chapter3/filter.def(2)"filter/a A p l. +img class="anchor" src="icons/tick.png" id="filter_true"lemma filter_true : ∀A,l,a,p. p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → + a href="cic:/matita/tutorial/chapter3/filter.def(2)"filter/a A p (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a a title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/a a href="cic:/matita/tutorial/chapter3/filter.def(2)"filter/a A p l. #A #l #a #p #pa (elim l) normalize >pa // qed. -lemma filter_false : ∀A,l,a,p. p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → - a href="cic:/matita/tutorial/chapter3/filter.def(2)"filter/a A p (aa title="cons" href="cic:/fakeuri.def(1)":/a:l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter3/filter.def(2)"filter/a A p l. +img class="anchor" src="icons/tick.png" id="filter_false"lemma filter_false : ∀A,l,a,p. p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → + a href="cic:/matita/tutorial/chapter3/filter.def(2)"filter/a A p (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter3/filter.def(2)"filter/a A p l. #A #l #a #p #pa (elim l) normalize >pa normalize // qed. (* As another example, let us redefine the map function using foldr. The @@ -227,7 +227,7 @@ result type B is (list B), the base value b is [], and the fold function of type A → list B → list B is the function mapping a and l to (f a)::l. *) -definition map_again ≝ λA,B,f,l. a href="cic:/matita/tutorial/chapter3/foldr.fix(0,4,1)"foldr/a A (a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a B) (λa,l.f aa title="cons" href="cic:/fakeuri.def(1)":/a:l) a title="nil" href="cic:/fakeuri.def(1)"[/a] l. +img class="anchor" src="icons/tick.png" id="map_again"definition map_again ≝ λA,B,f,l. a href="cic:/matita/tutorial/chapter3/foldr.fix(0,4,1)"foldr/a A (a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a B) (λa,l.f aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al) a title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a l. (* h2 class="section"Extensional equality/h2 @@ -240,19 +240,19 @@ and they are clearly different. What we would like to say is that the two programs behave in the same way: this is a different, extensional equality that can be defined in the following way. *) -definition ExtEq ≝ λA,B:Type[0].λf,g:A→B.∀a:A.f a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a g a. +img class="anchor" src="icons/tick.png" id="ExtEq"definition ExtEq ≝ λA,B:Type[0].λf,g:A→B.∀a:A.f a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a g a. (* Proving that map and map_again are extentionally equal in the previous sense can be proved by a trivial structural induction on the list *) -lemma eq_maps: ∀A,B,f. a href="cic:/matita/tutorial/chapter3/ExtEq.def(1)"ExtEq/a ?? (a href="cic:/matita/tutorial/chapter3/map.fix(0,3,1)"map/a A B f) (a href="cic:/matita/tutorial/chapter3/map_again.def(2)"map_again/a A B f). +img class="anchor" src="icons/tick.png" id="eq_maps"lemma eq_maps: ∀A,B,f. a href="cic:/matita/tutorial/chapter3/ExtEq.def(1)"ExtEq/a ?? (a href="cic:/matita/tutorial/chapter3/map.fix(0,3,1)"map/a A B f) (a href="cic:/matita/tutorial/chapter3/map_again.def(2)"map_again/a A B f). #A #B #f #n (elim n) normalize // qed. (* Let us make another remark about extensional equality. It is clear that, if f is extensionally equal to g, then (map A B f) is extensionally equal to (map A B g). Let us prove it. *) -theorem eq_map : ∀A,B,f,g. a href="cic:/matita/tutorial/chapter3/ExtEq.def(1)"ExtEq/a A B f g → a href="cic:/matita/tutorial/chapter3/ExtEq.def(1)"ExtEq/a ?? (a href="cic:/matita/tutorial/chapter3/map.fix(0,3,1)"map/a span style="text-decoration: underline;"/spanA B f) (a href="cic:/matita/tutorial/chapter3/map.fix(0,3,1)"map/a A B g). +img class="anchor" src="icons/tick.png" id="eq_map"theorem eq_map : ∀A,B,f,g. a href="cic:/matita/tutorial/chapter3/ExtEq.def(1)"ExtEq/a A B f g → a href="cic:/matita/tutorial/chapter3/ExtEq.def(1)"ExtEq/a ?? (a href="cic:/matita/tutorial/chapter3/map.fix(0,3,1)"map/a span style="text-decoration: underline;"/spanA B f) (a href="cic:/matita/tutorial/chapter3/map.fix(0,3,1)"map/a A B g). #A #B #f #g #eqfg (* the relevant point is that we cannot proceed by rewriting f with g via @@ -275,7 +275,7 @@ A really convenient tool is the following combination of fold and filter, that essentially allow you to iterate on every subset of a given enumerated (finite) type, represented as a list. *) - let rec fold (A,B:Type[0]) (op:B→B→B) (b:B) (p:A→a href="cic:/matita/basics/bool/bool.ind(1,0,0)" title="null"bool/a) (f:A→B) (l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) on l:B ≝ + img class="anchor" src="icons/tick.png" id="fold"let rec fold (A,B:Type[0]) (op:B→B→B) (b:B) (p:A→a href="cic:/matita/basics/bool/bool.ind(1,0,0)" title="null"bool/a) (f:A→B) (l:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A) on l:B ≝ match l with [ nil ⇒ b | cons a l ⇒ if p a then op (f a) (fold A B op b p f l) else @@ -294,38 +294,38 @@ for @{'fold $op $nil (λ${ident i}.true) (λ${ident i}. $f) $l}. interpretation "\fold" 'fold op nil p f l = (fold ? ? op nil p f l). -theorem fold_true: +img class="anchor" src="icons/tick.png" id="fold_true"theorem fold_true: ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f:A→B. p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → - a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ aa title="cons" href="cic:/fakeuri.def(1)":/a:l| p i} (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a - op (f a) a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ l| p i} (f i). + a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a + op (f a) a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ l| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). #A #B #a #l #p #op #nil #f #pa normalize >pa // qed. -theorem fold_false: +img class="anchor" src="icons/tick.png" id="fold_false"theorem fold_false: ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f. -p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ aa title="cons" href="cic:/fakeuri.def(1)":/a:l| p i} (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a - a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ l| p i} (f i). +p a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a + a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ l| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). #A #B #a #l #p #op #nil #f #pa normalize >pa // qed. -theorem fold_filter: +img class="anchor" src="icons/tick.png" id="fold_filter"theorem fold_filter: ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f:A →B. - a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ l| p i} (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a - a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ (a href="cic:/matita/tutorial/chapter3/filter.def(2)"filter/a A p l)} (f i). + a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ l| p ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a + a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ (a href="cic:/matita/tutorial/chapter3/filter.def(2)"filter/a A p l)a title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). #A #B #a #l #p #op #nil #f elim l // #a #tl #Hind cases(a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (p a)) #pa [ >a href="cic:/matita/tutorial/chapter3/filter_true.def(3)"filter_true/a // > a href="cic:/matita/tutorial/chapter3/fold_true.def(3)"fold_true/a // >a href="cic:/matita/tutorial/chapter3/fold_true.def(3)"fold_true/a // | >a href="cic:/matita/tutorial/chapter3/filter_false.def(3)"filter_false/a // >a href="cic:/matita/tutorial/chapter3/fold_false.def(3)"fold_false/a // ] qed. -record Aop (A:Type[0]) (nil:A) : Type[0] ≝ +img class="anchor" src="icons/tick.png" id="Aop"record Aop (A:Type[0]) (nil:A) : Type[0] ≝ {op :2> A → A → A; nill:∀a. op nil a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a; nilr:∀a. op a nil a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a; assoc: ∀a,b,c.op a (op b c) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a op (op a b) c }. -theorem fold_sum: ∀A,B. ∀I,J:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A.∀nil.∀op:a href="cic:/matita/tutorial/chapter3/Aop.ind(1,0,2)"Aop/a B nil.∀f:A → B. - op (a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ I} (f i)) (a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ J} (f i)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a - a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ (Ia title="append" href="cic:/fakeuri.def(1)"@/aJ)} (f i). +img class="anchor" src="icons/tick.png" id="fold_sum"theorem fold_sum: ∀A,B. ∀I,J:a href="cic:/matita/tutorial/chapter3/list.ind(1,0,1)"list/a A.∀nil.∀op:a href="cic:/matita/tutorial/chapter3/Aop.ind(1,0,2)"Aop/a B nil.∀f:A → B. + op (a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ Ia title="\fold" href="cic:/fakeuri.def(1)"}/a (f i)) (a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ Ja title="\fold" href="cic:/fakeuri.def(1)"}/a (f i)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a + a title="\fold" href="cic:/fakeuri.def(1)"\fold/a[op,nil]_{i ∈ (Ia title="append" href="cic:/fakeuri.def(1)"@/aJ)a title="\fold" href="cic:/fakeuri.def(1)"}/a (f i). #A #B #I #J #nil #op #f (elim I) normalize [>a href="cic:/matita/tutorial/chapter3/nill.fix(0,2,2)"nill/a//|#a #tl #Hind <a href="cic:/matita/tutorial/chapter3/assoc.fix(0,2,2)"assoc/a //] qed. \ No newline at end of file diff --git a/weblib/tutorial/chapter4.ma b/weblib/tutorial/chapter4.ma index ff6071764..7f28c5ddb 100644 --- a/weblib/tutorial/chapter4.ma +++ b/weblib/tutorial/chapter4.ma @@ -9,14 +9,14 @@ characteristic predicates over some universe codeA/code, that is as objects A→Prop. For instance the empty set is defined by the always false function: *) -definition empty_set ≝ λA:Type[0].λa:A.a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a. +img class="anchor" src="icons/tick.png" id="empty_set"definition empty_set ≝ λA:Type[0].λa:A.a href="cic:/matita/basics/logic/False.ind(1,0,0)"False/a. notation "\emptyv" non associative with precedence 90 for @{'empty_set}. interpretation "empty set" 'empty_set = (empty_set ?). (* Similarly, a singleton set contaning containing an element a, is defined by by the characteristic function asserting equality with a *) -definition singleton ≝ λA.λx,a:A.xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aspan class="error" title="Parse error: [term] expected after [sym=] (in [term])"/spana. +img class="anchor" src="icons/tick.png" id="singleton"definition singleton ≝ λA.λx,a:A.xa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aspan class="error" title="Parse error: [term] expected after [sym=] (in [term])"/spana. (* notation "{x}" non associative with precedence 90 for @{'sing_lang $x}. *) interpretation "singleton" 'singl x = (singleton ? x). @@ -26,22 +26,22 @@ The operations of union, intersection, complement and substraction are easily defined in terms of the propositional connectives of dijunction, conjunction and negation *) -definition union : ∀A:Type[0].∀P,Q.A → Prop ≝ λA,P,Q,a.P a a title="logical or" href="cic:/fakeuri.def(1)"∨/a Q a. +img class="anchor" src="icons/tick.png" id="union"definition union : ∀A:Type[0].∀P,Q.A → Prop ≝ λA,P,Q,a.P a a title="logical or" href="cic:/fakeuri.def(1)"∨/a Q a. interpretation "union" 'union a b = (union ? a b). -definition intersection : ∀A:Type[0].∀P,Q.A→Prop ≝ λA,P,Q,a.P a a title="logical and" href="cic:/fakeuri.def(1)"∧/aspan class="error" title="Parse error: [term] expected after [sym∧] (in [term])"/span Q a. +img class="anchor" src="icons/tick.png" id="intersection"definition intersection : ∀A:Type[0].∀P,Q.A→Prop ≝ λA,P,Q,a.P a a title="logical and" href="cic:/fakeuri.def(1)"∧/aspan class="error" title="Parse error: [term] expected after [sym∧] (in [term])"/span Q a. interpretation "intersection" 'intersects a b = (intersection ? a b). -definition complement ≝ λU:Type[0].λA:U → Prop.λw.a title="logical not" href="cic:/fakeuri.def(1)"¬/a A w. +img class="anchor" src="icons/tick.png" id="complement"definition complement ≝ λU:Type[0].λA:U → Prop.λw.a title="logical not" href="cic:/fakeuri.def(1)"¬/a A w. interpretation "complement" 'not a = (complement ? a). -definition substraction := λU:Type[0].λA,B:U → Prop.λw.A w a title="logical and" href="cic:/fakeuri.def(1)"∧/a a title="logical not" href="cic:/fakeuri.def(1)"¬/a B w. +img class="anchor" src="icons/tick.png" id="substraction"definition substraction := λU:Type[0].λA,B:U → Prop.λw.A w a title="logical and" href="cic:/fakeuri.def(1)"∧/a a title="logical not" href="cic:/fakeuri.def(1)"¬/a B w. interpretation "substraction" 'minus a b = (substraction ? a b). (* Finally, we use implication to define the inclusion relation between sets *) -definition subset: ∀A:Type[0].∀P,Q:A→Prop.Prop ≝ λA,P,Q.∀a:A.(P a → Q a). +img class="anchor" src="icons/tick.png" id="subset"definition subset: ∀A:Type[0].∀P,Q:A→Prop.Prop ≝ λA,P,Q.∀a:A.(P a → Q a). interpretation "subset" 'subseteq a b = (subset ? a b). (* @@ -49,7 +49,7 @@ interpretation "subset" 'subseteq a b = (subset ? a b). Two sets are equals if and only if they have the same elements, that is, if the two characteristic functions are extensionally equivalent: *) -definition eqP ≝ λA:Type[0].λP,Q:A → Prop.∀a:A.P a a title="iff" href="cic:/fakeuri.def(1)"↔/aspan class="error" title="Parse error: [term] expected after [sym↔] (in [term])"/span Q a. +img class="anchor" src="icons/tick.png" id="eqP"definition eqP ≝ λA:Type[0].λP,Q:A → Prop.∀a:A.P a a title="iff" href="cic:/fakeuri.def(1)"↔/aspan class="error" title="Parse error: [term] expected after [sym↔] (in [term])"/span Q a. notation "A =1 B" non associative with precedence 45 for @{'eqP $A $B}. interpretation "extensional equality" 'eqP a b = (eqP ? a b). @@ -58,38 +58,38 @@ This notion of equality is different from the intensional equality of functions; the fact it defines an equivalence relation must be explicitly proved: *) -lemma eqP_sym: ∀U.∀A,B:U →Prop. +img class="anchor" src="icons/tick.png" id="eqP_sym"lemma eqP_sym: ∀U.∀A,B:U →Prop. A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 B → B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A. #U #A #B #eqAB #a @a href="cic:/matita/basics/logic/iff_sym.def(2)"iff_sym/a @eqAB qed. -lemma eqP_trans: ∀U.∀A,B,C:U →Prop. +img class="anchor" src="icons/tick.png" id="eqP_trans"lemma eqP_trans: ∀U.∀A,B,C:U →Prop. A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 B → B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C → A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C. #U #A #B #C #eqAB #eqBC #a @a href="cic:/matita/basics/logic/iff_trans.def(2)"iff_trans/a // qed. (* For the same reason, we must also prove that all the operations behave well with respect to eqP: *) -lemma eqP_union_r: ∀U.∀A,B,C:U →Prop. +img class="anchor" src="icons/tick.png" id="eqP_union_r"lemma eqP_union_r: ∀U.∀A,B,C:U →Prop. A a title="extensional equality" href="cic:/fakeuri.def(1)"=/aspan class="error" title="Parse error: NUMBER '1' or [term] expected after [sym=] (in [term])"/span1 C → A a title="union" href="cic:/fakeuri.def(1)"∪/a B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C a title="union" href="cic:/fakeuri.def(1)"∪/a B. #U #A #B #C #eqAB #a @a href="cic:/matita/basics/logic/iff_or_r.def(2)"iff_or_r/a @eqAB qed. -lemma eqP_union_l: ∀U.∀A,B,C:U →Prop. +img class="anchor" src="icons/tick.png" id="eqP_union_l"lemma eqP_union_l: ∀U.∀A,B,C:U →Prop. B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C → A a title="union" href="cic:/fakeuri.def(1)"∪/a B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A a title="union" href="cic:/fakeuri.def(1)"∪/a C. #U #A #B #C #eqBC #a @a href="cic:/matita/basics/logic/iff_or_l.def(2)"iff_or_l/a @eqBC qed. -lemma eqP_intersect_r: ∀U.∀A,B,C:U →Prop. +img class="anchor" src="icons/tick.png" id="eqP_intersect_r"lemma eqP_intersect_r: ∀U.∀A,B,C:U →Prop. A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C → A a title="intersection" href="cic:/fakeuri.def(1)"∩/a B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C a title="intersection" href="cic:/fakeuri.def(1)"∩/a B. #U #A #B #C #eqAB #a @a href="cic:/matita/basics/logic/iff_and_r.def(2)"iff_and_r/a @eqAB qed. -lemma eqP_intersect_l: ∀U.∀A,B,C:U →Prop. +img class="anchor" src="icons/tick.png" id="eqP_intersect_l"lemma eqP_intersect_l: ∀U.∀A,B,C:U →Prop. B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C → A a title="intersection" href="cic:/fakeuri.def(1)"∩/aspan class="error" title="Parse error: [term] expected after [sym∩] (in [term])"/span B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A a title="intersection" href="cic:/fakeuri.def(1)"∩/a C. #U #A #B #C #eqBC #a @a href="cic:/matita/basics/logic/iff_and_l.def(2)"iff_and_l/a @eqBC qed. -lemma eqP_substract_r: ∀U.∀A,B,C:U →Prop. +img class="anchor" src="icons/tick.png" id="eqP_substract_r"lemma eqP_substract_r: ∀U.∀A,B,C:U →Prop. A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C → A a title="substraction" href="cic:/fakeuri.def(1)"-/a B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C a title="substraction" href="cic:/fakeuri.def(1)"-/a B. #U #A #B #C #eqAB #a @a href="cic:/matita/basics/logic/iff_and_r.def(2)"iff_and_r/a @eqAB qed. -lemma eqP_substract_l: ∀U.∀A,B,C:U →Prop. +img class="anchor" src="icons/tick.png" id="eqP_substract_l"lemma eqP_substract_l: ∀U.∀A,B,C:U →Prop. B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C → A a title="substraction" href="cic:/fakeuri.def(1)"-/a B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A a title="substraction" href="cic:/fakeuri.def(1)"-/a C. #U #A #B #C #eqBC #a @a href="cic:/matita/basics/logic/iff_and_l.def(2)"iff_and_l/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/iff_not.def(4)"iff_not/a/span/span/ qed. @@ -99,16 +99,16 @@ We can now prove several properties of the previous set-theoretic operations. In particular, union is commutative and associative, and the empty set is an identity element: *) -lemma union_empty_r: ∀U.∀A:U→Prop. +img class="anchor" src="icons/tick.png" id="union_empty_r"lemma union_empty_r: ∀U.∀A:U→Prop. A a title="union" href="cic:/fakeuri.def(1)"∪/a a title="empty set" href="cic:/fakeuri.def(1)"∅/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A. #U #A #w % [* // normalize #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace /span/span/ | /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a/span/span/] qed. -lemma union_comm : ∀U.∀A,B:U →Prop. +img class="anchor" src="icons/tick.png" id="union_comm"lemma union_comm : ∀U.∀A,B:U →Prop. A a title="union" href="cic:/fakeuri.def(1)"∪/a B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 B a title="union" href="cic:/fakeuri.def(1)"∪/a A. #U #A #B #a % * /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ qed. -lemma union_assoc: ∀U.∀A,B,C:U → Prop. +img class="anchor" src="icons/tick.png" id="union_assoc"lemma union_assoc: ∀U.∀A,B,C:U → Prop. A a title="union" href="cic:/fakeuri.def(1)"∪/a B a title="union" href="cic:/fakeuri.def(1)"∪/a C a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A a title="union" href="cic:/fakeuri.def(1)"∪/a (B a title="union" href="cic:/fakeuri.def(1)"∪/a C). #S #A #B #C #w % [* [* /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ | /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ ] | * [/span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a/span/span/ | * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/] qed. @@ -116,39 +116,39 @@ qed. (* In the same way we prove commutativity and associativity for set interesection *) -lemma cap_comm : ∀U.∀A,B:U →Prop. +img class="anchor" src="icons/tick.png" id="cap_comm"lemma cap_comm : ∀U.∀A,B:U →Prop. A a title="intersection" href="cic:/fakeuri.def(1)"∩/a B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 B a title="intersection" href="cic:/fakeuri.def(1)"∩/a A. #U #A #B #a % * /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ qed. -lemma cap_assoc: ∀U.∀A,B,C:U→Prop. +img class="anchor" src="icons/tick.png" id="cap_assoc"lemma cap_assoc: ∀U.∀A,B,C:U→Prop. A a title="intersection" href="cic:/fakeuri.def(1)"∩/a (B a title="intersection" href="cic:/fakeuri.def(1)"∩/a C) a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 (A a title="intersection" href="cic:/fakeuri.def(1)"∩/a B) a title="intersection" href="cic:/fakeuri.def(1)"∩/a C. #U #A #B #C #w % [ * #Aw * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ span class="autotactic"span class="autotrace"/span/span| * * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ ] qed. (* We can also easily prove idempotency for union and intersection *) -lemma union_idemp: ∀U.∀A:U →Prop. +img class="anchor" src="icons/tick.png" id="union_idemp"lemma union_idemp: ∀U.∀A:U →Prop. A a title="union" href="cic:/fakeuri.def(1)"∪/a A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A. #U #A #a % [* // | /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/] qed. -lemma cap_idemp: ∀U.∀A:U →Prop. +img class="anchor" src="icons/tick.png" id="cap_idemp"lemma cap_idemp: ∀U.∀A:U →Prop. A a title="intersection" href="cic:/fakeuri.def(1)"∩/a A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A. #U #A #a % [* // | /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/] qed. (* We conclude our examples with a couple of distributivity theorems, and a characterization of substraction in terms of interesection and complementation. *) -lemma distribute_intersect : ∀U.∀A,B,C:U→Prop. +img class="anchor" src="icons/tick.png" id="distribute_intersect"lemma distribute_intersect : ∀U.∀A,B,C:U→Prop. (A a title="union" href="cic:/fakeuri.def(1)"∪/a B) a title="intersection" href="cic:/fakeuri.def(1)"∩/a C a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 (A a title="intersection" href="cic:/fakeuri.def(1)"∩/a C) a title="union" href="cic:/fakeuri.def(1)"∪/a (B a title="intersection" href="cic:/fakeuri.def(1)"∩/a C). #U #A #B #C #w % [* * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a, a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ | * * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a, a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/] qed. -lemma distribute_substract : ∀U.∀A,B,C:U→Prop. +img class="anchor" src="icons/tick.png" id="distribute_substract"lemma distribute_substract : ∀U.∀A,B,C:U→Prop. (A a title="union" href="cic:/fakeuri.def(1)"∪/a B) a title="substraction" href="cic:/fakeuri.def(1)"-/a C a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 (A a title="substraction" href="cic:/fakeuri.def(1)"-/a C) a title="union" href="cic:/fakeuri.def(1)"∪/a (B a title="substraction" href="cic:/fakeuri.def(1)"-/a C). #U #A #B #C #w % [* * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a, a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ | * * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a, a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/] qed. -lemma substract_def:∀U.∀A,B:U→Prop. Aa title="substraction" href="cic:/fakeuri.def(1)"-/aB a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A a title="intersection" href="cic:/fakeuri.def(1)"∩/a a title="complement" href="cic:/fakeuri.def(1)"¬/aB. +img class="anchor" src="icons/tick.png" id="substract_def"lemma substract_def:∀U.∀A,B:U→Prop. Aa title="substraction" href="cic:/fakeuri.def(1)"-/aB a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A a title="intersection" href="cic:/fakeuri.def(1)"∩/a a title="complement" href="cic:/fakeuri.def(1)"¬/aB. #U #A #B #w normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ qed. @@ -159,7 +159,7 @@ between elements of a set U, namely a boolean function eqb: U→U→bool such th for any pair of elements a and b in U, (eqb x y) is true if and only if x=y. A set equipped with such an equality is called a DeqSet: *) -record DeqSet : Type[1] ≝ { carr :> Type[0]; +img class="anchor" src="icons/tick.png" id="DeqSet"record DeqSet : Type[1] ≝ { carr :> Type[0]; eqb: carr → carr → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a; eqb_true: ∀x,y. (eqb x y a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) a title="iff" href="cic:/fakeuri.def(1)"↔/a (x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a y) }. @@ -188,7 +188,7 @@ notation "\b H" non associative with precedence 90 statement asserts that we can reflect a proof that eqb a b is false into a proof of the proposition a ≠ b. *) -lemma eqb_false: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a,b:S. +img class="anchor" src="icons/tick.png" id="eqb_false"lemma eqb_false: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a,b:S. (a href="cic:/matita/tutorial/chapter4/eqb.fix(0,0,3)"eqb/a ? a b) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a a title="iff" href="cic:/fakeuri.def(1)"↔/a a a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a b. (* We start the proof introducing the hypothesis, and then split the "if" and @@ -225,7 +225,7 @@ notation "\bf H" non associative with precedence 90 (* The following statement proves that propositional equality in a DeqSet is decidable in the traditional sense, namely either a=b or a≠b *) - lemma dec_eq: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a,b:S. a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b a title="logical or" href="cic:/fakeuri.def(1)"∨/a a a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a b. + img class="anchor" src="icons/tick.png" id="dec_eq"lemma dec_eq: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a,b:S. a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b a title="logical or" href="cic:/fakeuri.def(1)"∨/a a a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a b. #S #a #b cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (a href="cic:/matita/tutorial/chapter4/eqb.fix(0,0,3)"eqb/a ? a b)) #H [%1 @(\P H) | %2 @(\Pf H)] qed. @@ -237,16 +237,16 @@ the boolean equality beqb, that is just the xand function, then prove that beqb b1 b2 is true if and only if b1=b2, and finally build the type DeqBool by instantiating the DeqSet record with the previous information *) -definition beqb ≝ λb1,b2. +img class="anchor" src="icons/tick.png" id="beqb"definition beqb ≝ λb1,b2. match b1 with [ true ⇒ b2 | false ⇒ a href="cic:/matita/basics/bool/notb.def(1)"notb/a b2]. notation < "a == b" non associative with precedence 45 for @{beqb $a $b }. -lemma beqb_true: ∀b1,b2. a href="cic:/matita/basics/logic/iff.def(1)"iff/a (a href="cic:/matita/tutorial/chapter4/beqb.def(2)"beqb/a b1 b2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) (b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b2). +img class="anchor" src="icons/tick.png" id="beqb_true"lemma beqb_true: ∀b1,b2. a href="cic:/matita/basics/logic/iff.def(1)"iff/a (a href="cic:/matita/tutorial/chapter4/beqb.def(2)"beqb/a b1 b2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) (b1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b2). #b1 #b2 cases b1 cases b2 normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ qed. -definition DeqBool ≝ a href="cic:/matita/tutorial/chapter4/DeqSet.con(0,1,0)"mk_DeqSet/a a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a a href="cic:/matita/tutorial/chapter4/beqb.def(2)"beqb/a a href="cic:/matita/tutorial/chapter4/beqb_true.def(4)"beqb_true/a. +img class="anchor" src="icons/tick.png" id="DeqBool"definition DeqBool ≝ a href="cic:/matita/tutorial/chapter4/DeqSet.con(0,1,0)"mk_DeqSet/a a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a a href="cic:/matita/tutorial/chapter4/beqb.def(2)"beqb/a a href="cic:/matita/tutorial/chapter4/beqb_true.def(4)"beqb_true/a. (* At this point, we would expect to be able to prove things like the following: for any boolean b, if b==false is true then b=false. @@ -254,7 +254,7 @@ Unfortunately, this would not work, unless we declare b of type DeqBool (change the type in the following statement and see what happens). *) -example exhint: ∀b:a href="cic:/matita/tutorial/chapter4/DeqBool.def(5)"DeqBool/a. (ba title="eqb" href="cic:/fakeuri.def(1)"=/a=a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → ba title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. +img class="anchor" src="icons/tick.png" id="exhint"example exhint: ∀b:a href="cic:/matita/tutorial/chapter4/DeqBool.def(5)"DeqBool/a. (ba title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → ba title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. #b #H @(\P H) qed. @@ -287,7 +287,7 @@ unification hint 0 a href="cic:/fakeuri.def(1)" title="hint_decl_Type0"≔/a (* After having provided the previous hints, we may rewrite example exhint declaring b of type bool. *) -example exhint1: ∀b:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. (b a title="eqb" href="cic:/fakeuri.def(1)"=/a= a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → b a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. +img class="anchor" src="icons/tick.png" id="exhint1"example exhint1: ∀b:a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a. (b a title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → b a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. #b #H @(\P H) qed. @@ -295,10 +295,10 @@ qed. this, we must as usual define the boolen equality function, and prove it correctly reflects propositional equality. *) -definition eq_pairs ≝ - λA,B:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λp1,p2:Aa title="Product" href="cic:/fakeuri.def(1)"×/aB.(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p1 a title="eqb" href="cic:/fakeuri.def(1)"=/a= a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p2) a title="boolean and" href="cic:/fakeuri.def(1)"∧/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p1 a title="eqb" href="cic:/fakeuri.def(1)"=/a= a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p2). +img class="anchor" src="icons/tick.png" id="eq_pairs"definition eq_pairs ≝ + λA,B:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λp1,p2:Aa title="Product" href="cic:/fakeuri.def(1)"×/aB.(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p1 a title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a p2) a title="boolean and" href="cic:/fakeuri.def(1)"∧/a (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p1 a title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/a a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a p2). -lemma eq_pairs_true: ∀A,B:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀p1,p2:Aa title="Product" href="cic:/fakeuri.def(1)"×/aB. +img class="anchor" src="icons/tick.png" id="eq_pairs_true"lemma eq_pairs_true: ∀A,B:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀p1,p2:Aa title="Product" href="cic:/fakeuri.def(1)"×/aB. a href="cic:/matita/tutorial/chapter4/eq_pairs.def(4)"eq_pairs/a A B p1 p2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="iff" href="cic:/fakeuri.def(1)"↔/a p1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a p2. #A #B * #a1 #b1 * #a2 #b2 % [#H cases (a href="cic:/matita/basics/bool/andb_true.def(5)"andb_true/a …H) normalize #eqa #eqb >(\P eqa) >(\P eqb) // @@ -306,7 +306,7 @@ lemma eq_pairs_true: ∀A,B:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1, ] qed. -definition DeqProd ≝ λA,B:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a. +img class="anchor" src="icons/tick.png" id="DeqProd"definition DeqProd ≝ λA,B:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a. a href="cic:/matita/tutorial/chapter4/DeqSet.con(0,1,0)"mk_DeqSet/a (Aa title="Product" href="cic:/fakeuri.def(1)"×/aB) (a href="cic:/matita/tutorial/chapter4/eq_pairs.def(4)"eq_pairs/a A B) (a href="cic:/matita/tutorial/chapter4/eq_pairs_true.def(6)"eq_pairs_true/a A B). (* Having an unification problem of the kind T1×T2 = carr X, what kind @@ -326,6 +326,6 @@ unification hint 0 a href="cic:/fakeuri.def(1)" title="hint_decl_Type0"≔/a (* ---------------------------------------- *) ⊢ a href="cic:/matita/tutorial/chapter4/eq_pairs.def(4)"eq_pairs/a T1 T2 p1 p2 ≡ a href="cic:/matita/tutorial/chapter4/eqb.fix(0,0,3)"eqb/a X p1 p2. -example hint2: ∀b1,b2. - a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ab1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉a title="eqb" href="cic:/fakeuri.def(1)"=/a=a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,b2〉a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ab1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,b2〉. +img class="anchor" src="icons/tick.png" id="hint2"example hint2: ∀b1,b2. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ab1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,b2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ab1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a,b2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. #b1 #b2 #H @(\P H). \ No newline at end of file diff --git a/weblib/tutorial/chapter5.ma b/weblib/tutorial/chapter5.ma index e5d226c32..f4594bc3b 100644 --- a/weblib/tutorial/chapter5.ma +++ b/weblib/tutorial/chapter5.ma @@ -13,10 +13,10 @@ include "tutorial/chapter4.ma". between an element x and a list l. Its definition is a straightforward recursion on l.*) -let rec memb (S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (x:S) (l: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/aspan class="error" title="Parse error: RPAREN expected after [term] (in [arg])"/span S) on l ≝ +img class="anchor" src="icons/tick.png" id="memb"let rec memb (S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (x:S) (l: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/aspan class="error" title="Parse error: RPAREN expected after [term] (in [arg])"/span S) on l ≝ match l with [ nil ⇒ a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a - | cons a tl ⇒ (x a title="eqb" href="cic:/fakeuri.def(1)"=/a= a) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a memb S x tl + | cons a tl ⇒ (x a title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/a a) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a memb S x tl ]span class="error" title="Parse error: NUMBER '1' or [term] or [sym=] expected after [sym=] (in [term])"/spanspan class="error" title="No choices for ID nil"/span. notation < "\memb x l" non associative with precedence 90 for @{'memb $x $l}. @@ -37,65 +37,65 @@ interpretation "boolean membership" 'memb a l = (memb ? a l). (op a b) is a member of (compose op l1 l2) *) -lemma memb_hd: ∀S,a,l. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (aa title="cons" href="cic:/fakeuri.def(1)":/a:l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +img class="anchor" src="icons/tick.png" id="memb_hd"lemma memb_hd: ∀S,a,l. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S #a #l normalize >(a href="cic:/matita/basics/logic/proj2.def(2)"proj2/a … (a href="cic:/matita/tutorial/chapter4/eqb_true.fix(0,0,4)"eqb_true/a S …) (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a S a)) // qed. -lemma memb_cons: ∀S,a,b,l. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/aspan class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"/span S a (ba title="cons" href="cic:/fakeuri.def(1)":/a:l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. -#S #a #b #l normalize cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b) normalize // +img class="anchor" src="icons/tick.png" id="memb_cons"lemma memb_cons: ∀S,a,b,l. + a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/aspan class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"/span S a (ba title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +#S #a #b #l normalize cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ab) normalize // qed. -lemma memb_single: ∀S,a,x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (xa title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a]) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x. -#S #a #x normalize cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x)) #H +img class="anchor" src="icons/tick.png" id="memb_single"lemma memb_single: ∀S,a,x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (xa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a x. +#S #a #x normalize cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ax)) #H [#_ >(\P H) // |>H normalize #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/] qed. -lemma memb_append: ∀S,a,l1,l2. +img class="anchor" src="icons/tick.png" id="memb_append"lemma memb_append: ∀S,a,l1,l2. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (l1a title="append" href="cic:/fakeuri.def(1)"@/aspan class="error" title="Parse error: [term level 46] expected after [sym@] (in [term])"/spanl2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="logical or" href="cic:/fakeuri.def(1)"∨/a a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S #a #l1span class="error" title="Parse error: illegal begin of statement"/spanspan class="error" title="Parse error: illegal begin of statement"/span elim l1 normalize [#l2 #H %2 //] -#b #tl #Hind #l2 cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b) normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a/span/span/ +#b #tl #Hind #l2 cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ab) normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a/span/span/ qed. -lemma memb_append_l1: ∀S,a,l1,l2. +img class="anchor" src="icons/tick.png" id="memb_append_l1"lemma memb_append_l1: ∀S,a,l1,l2. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (l1a title="append" href="cic:/fakeuri.def(1)"@/al2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S #a #l1 elim l1 normalize [normalize #le #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ - |#b #tl #Hind #l2 cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ + |#b #tl #Hind #l2 cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ab) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ ] qed. -lemma memb_append_l2: ∀S,a,l1,l2. +img class="anchor" src="icons/tick.png" id="memb_append_l2"lemma memb_append_l2: ∀S,a,l1,l2. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l2a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (l1a title="append" href="cic:/fakeuri.def(1)"@/al2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S #a #l1 elim l1 normalize // -#b #tl #Hind #l2 cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ +#b #tl #Hind #l2 cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ab) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ qed. -lemma memb_exists: ∀S,a,l.a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aspan class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"/span → a title="exists" href="cic:/fakeuri.def(1)"∃/al1,l2.la title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/al1a title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/a:l2). +img class="anchor" src="icons/tick.png" id="memb_exists"lemma memb_exists: ∀S,a,l.a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aspan class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"/span → a title="exists" href="cic:/fakeuri.def(1)"∃/al1,l2a title="exists" href="cic:/fakeuri.def(1)"./ala title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/al1a title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al2). #S #a #l elim l [normalize #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/] #b #tl #Hind #H cases (a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a … H) [#eqba @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … (a href="cic:/matita/basics/list/list.con(0,1,1)"nil/a S)) @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … tl) >(\P eqba) // |#mem_tl cases (Hind mem_tl) #l1 * #l2 #eqtl - @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … (ba title="cons" href="cic:/fakeuri.def(1)":/a:l1)) @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … l2) >eqtl // + @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … (ba title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al1)) @(a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a … l2) >eqtl // ] qed. -lemma not_memb_to_not_eq: ∀S,a,b,l. - a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S b l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. -#S #a #b #l cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b)) // +img class="anchor" src="icons/tick.png" id="not_memb_to_not_eq"lemma not_memb_to_not_eq: ∀S,a,b,l. + a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S b l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ab a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a. +#S #a #b #l cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ab)) // #eqab >(\P eqab) #H >H #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ qed. -lemma memb_map: ∀S1,S2,f,a,l. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S1 a la title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → +img class="anchor" src="icons/tick.png" id="memb_map"lemma memb_map: ∀S1,S2,f,a,l. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S1 a la title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S2 (f a) (a href="cic:/matita/basics/list/map.fix(0,3,1)"map/a … f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S1 #S2 #f #a #l elim l normalize [//] -#x #tl #memba cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x)) +#x #tl #memba cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ax)) [#eqx >eqx >(\P eqx) >(\b (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a … (f x))) normalize // - |#eqx >eqx cases (f aa title="eqb" href="cic:/fakeuri.def(1)"=/a=f x) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ + |#eqx >eqx cases (f aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/af x) normalize /span class="autotactic"2span class="autotrace" trace /span/span/ ] qed. -lemma memb_compose: ∀S1,S2,S3,op,a1,a2,l1,l2. +img class="anchor" src="icons/tick.png" id="memb_compose"lemma memb_compose: ∀S1,S2,S3,op,a1,a2,l1,l2. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S1 a1 l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S2 a2 l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S3 (op a1 a2) (a href="cic:/matita/basics/list/compose.def(2)"compose/a S1 S2 S3 op l1 l2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S1 #S2 #S3 #op #a1 #a2 #l1 elim l1 [normalize //] @@ -111,7 +111,7 @@ If we are interested in representing finite sets as lists, is is convenient to avoid duplications of elements. The following uniqueb check this property. *) -let rec uniqueb (S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) l on l : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ +img class="anchor" src="icons/tick.png" id="uniqueb"let rec uniqueb (S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) l on l : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ match l with [ nil ⇒ a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a | cons a tl ⇒ a href="cic:/matita/basics/bool/notb.def(1)"notb/a (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a tl) a title="boolean and" href="cic:/fakeuri.def(1)"∧/a uniqueb S tl @@ -119,29 +119,29 @@ let rec uniqueb (S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"Deq (* unique_append l1 l2 add l1 in fornt of l2, but preserving unicity *) -let rec unique_append (S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (l1,l2: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S) on l1 ≝ +img class="anchor" src="icons/tick.png" id="unique_append"let rec unique_append (S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (l1,l2: a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a S) on l1 ≝ match l1 with [ nil ⇒ l2 | cons a tl ⇒ let r ≝ unique_append S tl l2 in - if a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a r then r else aa title="cons" href="cic:/fakeuri.def(1)":/a:r + if a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a r then r else aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/ar ]. -lemma memb_unique_append: ∀S,a,l1,l2. +img class="anchor" src="icons/tick.png" id="memb_unique_append"lemma memb_unique_append: ∀S,a,l1,l2. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S l1 l2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l1a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="logical or" href="cic:/fakeuri.def(1)"∨/a a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S #a #l1 elim l1 normalize [#l2 #H %2 //] -#b #tl #Hind #l2 cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b)) #Hab >Hab normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a/span/span/ +#b #tl #Hind #l2 cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ab)) #Hab >Hab normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a/span/span/ cases (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S b (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S tl l2)) normalize [@Hind | >Hab normalize @Hind] qed. -lemma unique_append_elim: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀P: S → Prop.∀l1,l2. +img class="anchor" src="icons/tick.png" id="unique_append_elim"lemma unique_append_elim: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀P: S → Prop.∀l1,l2. (∀x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aspan class="error" title="Parse error: NUMBER '1' or [term] or [sym=] expected after [sym=] (in [term])"/span a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → P x) → (∀x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → P x) → ∀x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S l1 l2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → P x. #S #P #l1 #l2 #Hl1 #Hl2 #x #membx cases (a href="cic:/matita/tutorial/chapter5/memb_unique_append.def(6)"memb_unique_append/aspan class="error" title="No choices for ID memb_unique_append"/span … membx) /span class="autotactic"2span class="autotrace" trace /span/span/ qed. -lemma unique_append_unique: ∀S,l1,l2. a href="cic:/matita/tutorial/chapter5/uniqueb.fix(0,1,5)"uniqueb/a S l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → +img class="anchor" src="icons/tick.png" id="unique_append_unique"lemma unique_append_unique: ∀S,l1,l2. a href="cic:/matita/tutorial/chapter5/uniqueb.fix(0,1,5)"uniqueb/a S l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/uniqueb.fix(0,1,5)"uniqueb/a S (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S l1 l2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S #l1 elim l1 normalize // #a #tl #Hind #l2 #uniquel2 cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S tl l2))) @@ -151,14 +151,14 @@ qed. (* h2 class="section"Sublists/h2 *) -definition sublist ≝ +img class="anchor" src="icons/tick.png" id="sublist"definition sublist ≝ λS,l1,l2.∀a. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. -lemma sublist_length: ∀S,l1,l2. - a href="cic:/matita/tutorial/chapter5/uniqueb.fix(0,1,5)"uniqueb/a S l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l1 l2 → a title="norm" href="cic:/fakeuri.def(1)"|/al1| a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a title="norm" href="cic:/fakeuri.def(1)"|/al2|. +img class="anchor" src="icons/tick.png" id="sublist_length"lemma sublist_length: ∀S,l1,l2. + a href="cic:/matita/tutorial/chapter5/uniqueb.fix(0,1,5)"uniqueb/a S l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l1 l2 → a title="norm" href="cic:/fakeuri.def(1)"|/al1a title="norm" href="cic:/fakeuri.def(1)"|/a a title="natural 'less or equal to'" href="cic:/fakeuri.def(1)"≤/a a title="norm" href="cic:/fakeuri.def(1)"|/al2a title="norm" href="cic:/fakeuri.def(1)"|/a. #S #l1 elim l1 // #a #tl #Hind #l2 #unique #sub -cut (a title="exists" href="cic:/fakeuri.def(1)"∃/al3,l4.l2a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/al3a title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/a:l4)) [@a href="cic:/matita/tutorial/chapter5/memb_exists.def(5)"memb_exists/a @sub //] +cut (a title="exists" href="cic:/fakeuri.def(1)"∃/al3,l4a title="exists" href="cic:/fakeuri.def(1)"./al2a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/al3a title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al4)) [@a href="cic:/matita/tutorial/chapter5/memb_exists.def(5)"memb_exists/a @sub //] * #l3 * #l4 #eql2 >eql2 >a href="cic:/matita/basics/list/length_append.def(2)"length_append/a normalize applyS a href="cic:/matita/arithmetics/nat/le_S_S.def(2)"le_S_S/a <a href="cic:/matita/basics/list/length_append.def(2)"length_append/a @Hind [@(a href="cic:/matita/basics/bool/andb_true_r.def(4)"andb_true_r/a … unique)] >eql2 in sub; #sub #x #membx @@ -171,11 +171,11 @@ cases (a href="cic:/matita/tutorial/chapter5/memb_append.def(5)"memb_append/a ] qed. -lemma sublist_unique_append_l1: +img class="anchor" src="icons/tick.png" id="sublist_unique_append_l1"lemma sublist_unique_append_l1: ∀S,l1,l2. a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l1 (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S l1 l2). #S #l1 elim l1 normalize [#l2 #S #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/] #x #tl #Hind #l2 #a -normalize cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x)) #eqax >eqax +normalize cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a … (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ax)) #eqax >eqax [<(\P eqax) cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S tl l2))) [#H >H normalize // | #H >H normalize >(\b (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a … a)) //] |cases (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S tl l2)) normalize @@ -183,14 +183,14 @@ normalize cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_ ] qed. -lemma sublist_unique_append_l2: +img class="anchor" src="icons/tick.png" id="sublist_unique_append_l2"lemma sublist_unique_append_l2: ∀S,l1,l2. a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l2 (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S l1 l2). #S #l1 elim l1 [normalize //] #x #tl #Hind normalize #l2 #a cases (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a S tl l2)) normalize -[@Hind | cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x) normalize // @Hind] +[@Hind | cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ax) normalize // @Hind] qed. -lemma decidable_sublist:∀S,l1,l2. +img class="anchor" src="icons/tick.png" id="decidable_sublist"lemma decidable_sublist:∀S,l1,l2. (a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l1 l2) a title="logical or" href="cic:/fakeuri.def(1)"∨/a a title="logical not" href="cic:/fakeuri.def(1)"¬/a(a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S l1 l2). #S #l1 #l2 elim l1 [%1 #a normalize in ⊢ (%→?); #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ @@ -207,30 +207,30 @@ qed. (*h2 class="section"Filtering/h2*) -lemma memb_filter_true: ∀S,f,a,l. +img class="anchor" src="icons/tick.png" id="memb_filter_true"lemma memb_filter_true: ∀S,f,a,l. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/basics/list/filter.def(2)"filter/a S f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → f a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S #f #a #l elim l [normalize #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/] #b #tl #Hind cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (f b)) #H normalize >H normalize [2:@Hind] -cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b)) #eqab +cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ab)) #eqab [#_ >(\P eqab) // | >eqab normalize @Hind] qed. -lemma memb_filter_memb: ∀S,f,a,l. +img class="anchor" src="icons/tick.png" id="memb_filter_memb"lemma memb_filter_memb: ∀S,f,a,l. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/basics/list/filter.def(2)"filter/a S f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S #f #a #l elim l [normalize //] #b #tl #Hind normalize (cases (f b)) normalize -cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=b) normalize // @Hind +cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ab) normalize // @Hind qed. -lemma memb_filter: ∀S,f,l,x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → +img class="anchor" src="icons/tick.png" id="memb_filter"lemma memb_filter: ∀S,f,l,x. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="logical and" href="cic:/fakeuri.def(1)"∧/a (f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a). /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a, a href="cic:/matita/tutorial/chapter5/memb_filter_memb.def(5)"memb_filter_memb/a, a href="cic:/matita/tutorial/chapter5/memb_filter_true.def(5)"memb_filter_true/a/span/span/ qed. -lemma memb_filter_l: ∀S,f,x,l. (f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → +img class="anchor" src="icons/tick.png" id="memb_filter_l"lemma memb_filter_l: ∀S,f,x,l. (f x a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x l a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S x (a href="cic:/matita/basics/list/filter.def(2)"filter/a ? f l) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S #f #x #l #fx elim l normalize // -#b #tl #Hind cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (xa title="eqb" href="cic:/fakeuri.def(1)"=/a=b)) #eqxb +#b #tl #Hind cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (xa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ab)) #eqxb [<(\P eqxb) >(\b (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a … x)) >fx normalize >(\b (a href="cic:/matita/basics/logic/eq.con(0,1,2)"refl/a … x)) normalize // |>eqxb cases (f b) normalize [>eqxb normalize @Hind| @Hind] ] @@ -240,7 +240,7 @@ qed. h2 class="section"Exists/h2 *) -let rec exists (A:Type[0]) (p:A → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ +img class="anchor" src="icons/tick.png" id="exists"let rec exists (A:Type[0]) (p:A → a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a) (l:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A) on l : a href="cic:/matita/basics/bool/bool.ind(1,0,0)"bool/a ≝ match l with [ nil ⇒ a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a | cons h t ⇒ a href="cic:/matita/basics/bool/orb.def(1)"orb/a (p h) (exists A p t) diff --git a/weblib/tutorial/chapter6.ma b/weblib/tutorial/chapter6.ma index 679331198..dd6d7cc78 100644 --- a/weblib/tutorial/chapter6.ma +++ b/weblib/tutorial/chapter6.ma @@ -7,7 +7,7 @@ of words over a given alphabet, that we shall represent as a predicate over word include "tutorial/chapter5.ma". (* A word (or string) over an alphabet S is just a list of elements of S.*) -definition word ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.a href="cic:/matita/basics/list/list.ind(1,0,1)"list/aspan class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"/span S. +img class="anchor" src="icons/tick.png" id="word"definition word ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.a href="cic:/matita/basics/list/list.ind(1,0,1)"list/aspan class="error" title="Parse error: SYMBOL '.' expected after [grafite_ncommand] (in [executable])"/span S. (* For any alphabet there is only one word of length 0, the iempty word/i, which is denoted by ϵ .*) @@ -31,8 +31,8 @@ operations induced by string concatenation, and in particular the concatenation A · B of two languages A and B, the so called Kleene's star A* of A and the derivative of a language A w.r.t. a given character a. *) -definition cat : ∀S,l1,l2,w.Prop ≝ - λS.λl1,l2.λw:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S.a title="exists" href="cic:/fakeuri.def(1)"∃/aspan class="error" title="Parse error: [sym_] or [ident] expected after [sym∃] (in [term])"/spanw1,w2.w1 a title="append" href="cic:/fakeuri.def(1)"@/a w2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a w a title="logical and" href="cic:/fakeuri.def(1)"∧/a l1 w1 a title="logical and" href="cic:/fakeuri.def(1)"∧/aspan class="error" title="Parse error: [term] expected after [sym∧] (in [term])"/span l2 w2. +img class="anchor" src="icons/tick.png" id="cat"definition cat : ∀S,l1,l2,w.Prop ≝ + λS.λl1,l2.λw:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S.a title="exists" href="cic:/fakeuri.def(1)"∃/aspan class="error" title="Parse error: [sym_] or [ident] expected after [sym∃] (in [term])"/spanw1,w2a title="exists" href="cic:/fakeuri.def(1)"./aw1 a title="append" href="cic:/fakeuri.def(1)"@/a w2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a w a title="logical and" href="cic:/fakeuri.def(1)"∧/a l1 w1 a title="logical and" href="cic:/fakeuri.def(1)"∧/aspan class="error" title="Parse error: [term] expected after [sym∧] (in [term])"/span l2 w2. notation "a · b" non associative with precedence 60 for @{ 'middot $a $b}. interpretation "cat lang" 'middot a b = (cat ? a b). @@ -45,40 +45,40 @@ w1,w2,...wk all belonging to l, such that l = w1w2...wk. We need to define the latter operations. The following flatten function takes in input a list of words and concatenates them together. *) -let rec flatten (S : a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (l : a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a (a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S)) on l : a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S ≝ -match l with [ nil ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/a ] | cons w tl ⇒ w a title="append" href="cic:/fakeuri.def(1)"@/a flatten ? tl ]. +img class="anchor" src="icons/tick.png" id="flatten"let rec flatten (S : a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (l : a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a (a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S)) on l : a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S ≝ +match l with [ nil ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/a a title="nil" href="cic:/fakeuri.def(1)"]/a | cons w tl ⇒ w a title="append" href="cic:/fakeuri.def(1)"@/a flatten ? tl ]. (* Given a list of words l and a language r, (conjunct l r) is true if and only if all words in l are in r, that is for every w in l, r w holds. *) -let rec conjunct (S : a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (l : a href="cic:/matita/basics/list/list.ind(1,0,1)"list/aspan class="error" title="Parse error: RPAREN expected after [term] (in [arg])"/span (a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S)) (r : a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S → Prop) on l: Prop ≝ +img class="anchor" src="icons/tick.png" id="conjunct"let rec conjunct (S : a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (l : a href="cic:/matita/basics/list/list.ind(1,0,1)"list/aspan class="error" title="Parse error: RPAREN expected after [term] (in [arg])"/span (a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S)) (r : a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S → Prop) on l: Prop ≝ match l with [ nil ⇒ a href="cic:/matita/basics/logic/True.ind(1,0,0)"True/a | cons w tl ⇒ r w a title="logical and" href="cic:/fakeuri.def(1)"∧/a conjunct ? tl r ]. (* We are ready to give the formal definition of the Kleene's star of l: a word w belongs to l* is and only if there exists a list of strings lw such that (conjunct lw l) and l = flatten lw. *) -definition star ≝ λS.λl.λw:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S.a title="exists" href="cic:/fakeuri.def(1)"∃/alw.a href="cic:/matita/tutorial/chapter6/flatten.fix(0,1,4)"flatten/a ? lw a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a w a title="logical and" href="cic:/fakeuri.def(1)"∧/a a href="cic:/matita/tutorial/chapter6/conjunct.fix(0,1,4)"conjunct/a ? lw l. +img class="anchor" src="icons/tick.png" id="star"definition star ≝ λS.λl.λw:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S.a title="exists" href="cic:/fakeuri.def(1)"∃/alw.a href="cic:/matita/tutorial/chapter6/flatten.fix(0,1,4)"flatten/a ? lw a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a w a title="logical and" href="cic:/fakeuri.def(1)"∧/a a href="cic:/matita/tutorial/chapter6/conjunct.fix(0,1,4)"conjunct/a ? lw l. notation "a ^ *" non associative with precedence 90 for @{ 'star $a}. interpretation "star lang" 'star l = (star ? l). (* The derivative of a language A with respect to a character a is the set of all strings w such that aw is in A. *) -definition deriv ≝ λS.λA:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S → Prop.λa,w. A (aa title="cons" href="cic:/fakeuri.def(1)":/a:w). +img class="anchor" src="icons/tick.png" id="deriv"definition deriv ≝ λS.λA:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S → Prop.λa,w. A (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aw). (* h2 class="section"Language equalities/h2 Equality between languages is just the usual extensional equality between sets. The operation of concatenation behaves well with respect to this equality. *) -lemma cat_ext_l: ∀S.∀A,B,C:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S →Prop. +img class="anchor" src="icons/tick.png" id="cat_ext_l"lemma cat_ext_l: ∀S.∀A,B,C:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S →Prop. A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C → A a title="cat lang" href="cic:/fakeuri.def(1)"·/a B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C a title="cat lang" href="cic:/fakeuri.def(1)"·/a B. #S #A #B #C #H #w % * #w1 * #w2 * * #eqw #inw1 #inw2 cases (H w1) /span class="autotactic"6span class="autotrace" trace a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a, a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ qed. -lemma cat_ext_r: ∀S.∀A,B,C:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S →Prop. +img class="anchor" src="icons/tick.png" id="cat_ext_r"lemma cat_ext_r: ∀S.∀A,B,C:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S →Prop. B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 C → A a title="cat lang" href="cic:/fakeuri.def(1)"·/a B a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A a title="cat lang" href="cic:/fakeuri.def(1)"·/a C. #S #A #B #C #H #w % * #w1 * #w2 * * #eqw #inw1 #inw2 cases (H w2) /span class="autotactic"6span class="autotrace" trace a href="cic:/matita/basics/logic/ex.con(0,1,2)"ex_intro/a, a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ @@ -86,7 +86,7 @@ qed. (* Concatenating a language with the empty language results in the empty language. *) -lemma cat_empty_l: ∀S.∀A:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S→Prop. a title="empty set" href="cic:/fakeuri.def(1)"∅/a a title="cat lang" href="cic:/fakeuri.def(1)"·/a A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="empty set" href="cic:/fakeuri.def(1)"∅/a. +img class="anchor" src="icons/tick.png" id="cat_empty_l"lemma cat_empty_l: ∀S.∀A:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S→Prop. a title="empty set" href="cic:/fakeuri.def(1)"∅/a a title="cat lang" href="cic:/fakeuri.def(1)"·/a A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="empty set" href="cic:/fakeuri.def(1)"∅/a. #S #A #w % [|*] * #w1 * #w2 * * #_ * qed. @@ -94,16 +94,16 @@ qed. empty string, results in the language l; that is {ϵ} is a left and right unit with respect to concatenation. *) -lemma epsilon_cat_r: ∀S.∀A:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S →Prop. - A a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="singleton" href="cic:/fakeuri.def(1)"{/aa title="epsilon" href="cic:/fakeuri.def(1)"ϵ/a} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A. +img class="anchor" src="icons/tick.png" id="epsilon_cat_r"lemma epsilon_cat_r: ∀S.∀A:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S →Prop. + A a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="singleton" href="cic:/fakeuri.def(1)"{/aa title="epsilon" href="cic:/fakeuri.def(1)"ϵ/aa title="singleton" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 A. #S #A #w % [* #w1 * #w2 * * #eqw #inw1 normalize #eqw2 (a href="cic:/matita/tutorial/chapter7/append_eq_nil.def(4)"append_eq_nil/a …H…) /span class="autotactic"2span class="autotrace" trace /span/span/ @@ -280,22 +280,22 @@ lemma not_epsilon_lp : ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1, ] qed. -lemma epsilon_to_true : ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a title="epsilon" href="cic:/fakeuri.def(1)"ϵ/a a title="in_prl mem" href="cic:/fakeuri.def(1)"∈/a e → a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. +img class="anchor" src="icons/tick.png" id="epsilon_to_true"lemma epsilon_to_true : ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a title="epsilon" href="cic:/fakeuri.def(1)"ϵ/a a title="in_prl mem" href="cic:/fakeuri.def(1)"∈/a e → a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a. #S * #i #b cases b // normalize #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ qed. -lemma true_to_epsilon : ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a title="epsilon" href="cic:/fakeuri.def(1)"ϵ/a a title="in_prl mem" href="cic:/fakeuri.def(1)"∈/a e. +img class="anchor" src="icons/tick.png" id="true_to_epsilon"lemma true_to_epsilon : ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a title="epsilon" href="cic:/fakeuri.def(1)"ϵ/a a title="in_prl mem" href="cic:/fakeuri.def(1)"∈/a e. #S * #i #b #btrue normalize in btrue; >btrue %2 // qed. -lemma minus_eps_item: ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i}a title="substraction" href="cic:/fakeuri.def(1)"-/aa title="singleton" href="cic:/fakeuri.def(1)"{/aa title="nil" href="cic:/fakeuri.def(1)"[/a ]}. +img class="anchor" src="icons/tick.png" id="minus_eps_item"lemma minus_eps_item: ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{ia title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{ia title="in_pl" href="cic:/fakeuri.def(1)"}/aa title="substraction" href="cic:/fakeuri.def(1)"-/aa title="singleton" href="cic:/fakeuri.def(1)"{/aa title="nil" href="cic:/fakeuri.def(1)"[/a a title="nil" href="cic:/fakeuri.def(1)"]/aa title="singleton" href="cic:/fakeuri.def(1)"}/a. #S #i #w % [#H whd % // normalize @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … (a href="cic:/matita/tutorial/chapter7/not_epsilon_lp.def(8)"not_epsilon_lp/a …i)) // |* // ] qed. -lemma minus_eps_pre: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a title="in_pl" href="cic:/fakeuri.def(1)"\sem/aspan class="error" title="Parse error: [sym{] expected after [sym\sem ] (in [term])"/span{a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e}a title="substraction" href="cic:/fakeuri.def(1)"-/aa title="singleton" href="cic:/fakeuri.def(1)"{/aa title="nil" href="cic:/fakeuri.def(1)"[/a ]}. +img class="anchor" src="icons/tick.png" id="minus_eps_pre"lemma minus_eps_pre: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a title="in_pl" href="cic:/fakeuri.def(1)"\sem/aspan class="error" title="Parse error: [sym{] expected after [sym\sem ] (in [term])"/span{a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a ea title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{ea title="in_prl" href="cic:/fakeuri.def(1)"}/aa title="substraction" href="cic:/fakeuri.def(1)"-/aa title="singleton" href="cic:/fakeuri.def(1)"{/aa title="nil" href="cic:/fakeuri.def(1)"[/a a title="nil" href="cic:/fakeuri.def(1)"]/aa title="singleton" href="cic:/fakeuri.def(1)"}/a. #S * #i * [>a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a normalize in ⊢ (??%?); #w % [/span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a, a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a/span/span/ | * * // #H1 #H2 @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a @(a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a …H1 H2)] diff --git a/weblib/tutorial/chapter8.ma b/weblib/tutorial/chapter8.ma index 905e6604a..0e3293c1e 100644 --- a/weblib/tutorial/chapter8.ma +++ b/weblib/tutorial/chapter8.ma @@ -26,11 +26,11 @@ If we define then, we just have •(i1+i2) = •(i1)⊕ •(i2). *) -definition lo ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λa,b:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a a a title="pitem or" href="cic:/fakeuri.def(1)"+/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a b,a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a a a title="boolean or" href="cic:/fakeuri.def(1)"∨/a a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a b〉. +img class="anchor" src="icons/tick.png" id="lo"definition lo ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λa,b:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a a a title="pitem or" href="cic:/fakeuri.def(1)"+/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a b,a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a a a title="boolean or" href="cic:/fakeuri.def(1)"∨/a a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. notation "a ⊕ b" left associative with precedence 60 for @{'oplus $a $b}. interpretation "oplus" 'oplus a b = (lo ? a b). -lemma lo_def: ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀b1,b2. a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,b1〉a title="oplus" href="cic:/fakeuri.def(1)"⊕/aa title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai2,b2〉a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1a title="pitem or" href="cic:/fakeuri.def(1)"+/ai2,b1a title="boolean or" href="cic:/fakeuri.def(1)"∨/ab2〉. +img class="anchor" src="icons/tick.png" id="lo_def"lemma lo_def: ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀b1,b2. a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,b1a title="Pair construction" href="cic:/fakeuri.def(1)"〉/aa title="oplus" href="cic:/fakeuri.def(1)"⊕/aa title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai2,b2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/aa title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/aa title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1a title="pitem or" href="cic:/fakeuri.def(1)"+/ai2,b1a title="boolean or" href="cic:/fakeuri.def(1)"∨/ab2a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. // qed. (* @@ -47,20 +47,20 @@ In turn, ◃ says how to concatenate an item with a pre, that is however extreme Let us come to the formalized definitions: *) -definition pre_concat_r ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λi:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - match e with [ mk_Prod i1 b ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i1, b〉]. +img class="anchor" src="icons/tick.png" id="pre_concat_r"definition pre_concat_r ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λi:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + match e with [ mk_Prod i1 b ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i1, ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a]. notation "i ◃ e" left associative with precedence 60 for @{'lhd $i $e}. interpretation "pre_concat_r" 'lhd i e = (pre_concat_r ? i e). -lemma eq_to_ex_eq: ∀S.∀A,B:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S → Prop. +img class="anchor" src="icons/tick.png" id="eq_to_ex_eq"lemma eq_to_ex_eq: ∀S.∀A,B:a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S → Prop. A a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a B → A a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 B. #S #A #B #H >H #x % // qed. (* The behaviour of ◃ is summarized by the following, easy lemma: *) -lemma sem_pre_concat_r : ∀S,i.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{i a title="pre_concat_r" href="cic:/fakeuri.def(1)"◃/a e} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i} a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e}. +img class="anchor" src="icons/tick.png" id="sem_pre_concat_r"lemma sem_pre_concat_r : ∀S,i.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{i a title="pre_concat_r" href="cic:/fakeuri.def(1)"◃/a ea title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{ia title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a ea title="forget" href="cic:/fakeuri.def(1)"|/aa title="in_l" href="cic:/fakeuri.def(1)"}/a a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{ea title="in_prl" href="cic:/fakeuri.def(1)"}/a. #S #i * #i1 #b1 cases b1 [2: @a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a //] >a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a >a href="cic:/matita/tutorial/chapter7/sem_cat.def(8)"sem_cat/a >a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a /span class="autotactic"2span class="autotrace" trace /span/span/ qed. @@ -71,11 +71,11 @@ is to abstract one of the two functions with respect to the other. In particular we abstract pre_concat_l with respect to an input bcast function from items to pres. *) -definition pre_concat_l ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λbcast:∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S → a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.λe1:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.λi2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. +img class="anchor" src="icons/tick.png" id="pre_concat_l"definition pre_concat_l ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λbcast:∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S → a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.λe1:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.λi2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. match e1 with [ mk_Prod i1 b1 ⇒ match b1 with [ true ⇒ (i1 a title="pre_concat_r" href="cic:/fakeuri.def(1)"◃/a (bcast ? i2)) - | false ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 + | false ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a ] ]. @@ -86,59 +86,59 @@ interpretation "item-pre concat" 'tril op a b = (pre_concat_l ? op a b). notation "•" non associative with precedence 60 for @{eclose ?}. -let rec eclose (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (i: a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S) on i : a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S ≝ +img class="anchor" src="icons/tick.png" id="eclose"let rec eclose (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (i: a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S) on i : a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S ≝ match i with - [ pz ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"pz/a ?, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a 〉 - | pe ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem epsilon" href="cic:/fakeuri.def(1)"ϵ/a, a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a 〉 - | ps x ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem pp" href="cic:/fakeuri.def(1)"`/a.x, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 - | pp x ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem pp" href="cic:/fakeuri.def(1)"`/a.x, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a 〉 + [ pz ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"pz/a ?, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a + | pe ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem epsilon" href="cic:/fakeuri.def(1)"ϵ/a, a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a + | ps x ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem pp" href="cic:/fakeuri.def(1)"`/aa title="pitem pp" href="cic:/fakeuri.def(1)"./ax, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a + | pp x ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem pp" href="cic:/fakeuri.def(1)"`/aa title="pitem pp" href="cic:/fakeuri.def(1)"./ax, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a | po i1 i2 ⇒ •i1 a title="oplus" href="cic:/fakeuri.def(1)"⊕/a •i2 | pc i1 i2 ⇒ •i1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2 - | pk i ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (•i))a title="pitem star" href="cic:/fakeuri.def(1)"^/a*,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉]. + | pk i ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (•i))a title="pitem star" href="cic:/fakeuri.def(1)"^/aa title="pitem star" href="cic:/fakeuri.def(1)"*/a,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a]. notation "• x" non associative with precedence 60 for @{'eclose $x}. interpretation "eclose" 'eclose x = (eclose ? x). (* Here are a few simple properties of ▹ and •(-) *) -lemma pcl_true : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a i1 a title="pre_concat_r" href="cic:/fakeuri.def(1)"◃/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2). +img class="anchor" src="icons/tick.png" id="pcl_true"lemma pcl_true : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a i1 a title="pre_concat_r" href="cic:/fakeuri.def(1)"◃/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2). // qed. -lemma pcl_true_bis : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2)〉. +img class="anchor" src="icons/tick.png" id="pcl_true_bis"lemma pcl_true_bis : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2)a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. #S #i1 #i2 normalize cases (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2) // qed. -lemma pcl_false: ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉. +img class="anchor" src="icons/tick.png" id="pcl_false"lemma pcl_false: ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. // qed. -lemma eclose_plus: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. +img class="anchor" src="icons/tick.png" id="eclose_plus"lemma eclose_plus: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a title="eclose" href="cic:/fakeuri.def(1)"•/a(i1 a title="pitem or" href="cic:/fakeuri.def(1)"+/a i2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="eclose" href="cic:/fakeuri.def(1)"•/ai1 a title="oplus" href="cic:/fakeuri.def(1)"⊕/a a title="eclose" href="cic:/fakeuri.def(1)"•/ai2. // qed. -lemma eclose_dot: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. +img class="anchor" src="icons/tick.png" id="eclose_dot"lemma eclose_dot: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a title="eclose" href="cic:/fakeuri.def(1)"•/a(i1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="eclose" href="cic:/fakeuri.def(1)"•/ai1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2. // qed. -lemma eclose_star: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a title="eclose" href="cic:/fakeuri.def(1)"•/aia title="pitem star" href="cic:/fakeuri.def(1)"^/a* a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a(a title="eclose" href="cic:/fakeuri.def(1)"•/ai))a title="pitem star" href="cic:/fakeuri.def(1)"^/a*,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉. +img class="anchor" src="icons/tick.png" id="eclose_star"lemma eclose_star: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="eclose" href="cic:/fakeuri.def(1)"•/aia title="pitem star" href="cic:/fakeuri.def(1)"^/aa title="pitem star" href="cic:/fakeuri.def(1)"*/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a(a title="eclose" href="cic:/fakeuri.def(1)"•/ai))a title="pitem star" href="cic:/fakeuri.def(1)"^/aa title="pitem star" href="cic:/fakeuri.def(1)"*/a,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. // qed. (* The definition of •(-) (eclose) can then be lifted from items to pres in the obvious way. *) -definition lift ≝ λS.λf:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S →a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. +img class="anchor" src="icons/tick.png" id="lift"definition lift ≝ λS.λf:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S →a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. match e with - [ mk_Prod i b ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (f i), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (f i) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a b〉]. + [ mk_Prod i b ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (f i), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (f i) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a]. -definition preclose ≝ λS. a href="cic:/matita/tutorial/chapter8/lift.def(2)"lift/a S (a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"eclose/a S). +img class="anchor" src="icons/tick.png" id="preclose"definition preclose ≝ λS. a href="cic:/matita/tutorial/chapter8/lift.def(2)"lift/a S (a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"eclose/a S). interpretation "preclose" 'eclose x = (preclose ? x). (* Obviously, broadcasting does not change the carrier of the item, as it is easily proved by structural induction. *) -lemma erase_bull : ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai)| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/ai|. +img class="anchor" src="icons/tick.png" id="erase_bull"lemma erase_bull : ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai)a title="forget" href="cic:/fakeuri.def(1)"|/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aia title="forget" href="cic:/fakeuri.def(1)"|/a. #S #i elim i // [ #i1 #i2 #IH1 #IH2 >a href="cic:/matita/tutorial/chapter7/erase_dot.def(4)"erase_dot/a a href="cic:/matita/tutorial/chapter8/eclose_dot.def(5)"eclose_dot/a cases (a title="eclose" href="cic:/fakeuri.def(1)"•/ai1) #i11 #b1 cases b1 // a href="cic:/matita/tutorial/chapter8/pcl_true_bis.def(5)"pcl_true_bis/a // @@ -156,8 +156,8 @@ sem_bullet \sem{•i} =1 \sem{i} ∪ \sem{|i|} The proof of sem_oplus is straightforward. *) -lemma sem_oplus: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="oplus" href="cic:/fakeuri.def(1)"⊕/a e2} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2}. +img class="anchor" src="icons/tick.png" id="sem_oplus"lemma sem_oplus: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="oplus" href="cic:/fakeuri.def(1)"⊕/a e2a title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1a title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2a title="in_prl" href="cic:/fakeuri.def(1)"}/a. #S * #i1 #b1 * #i2 #b2 #w % [cases b1 cases b2 normalize /span class="autotactic"2span class="autotrace" trace /span/span/ * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ |cases b1 cases b2 normalize /span class="autotactic"2span class="autotrace" trace /span/span/ * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ * /span class="autotactic"3span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a, a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ @@ -174,13 +174,13 @@ sem_pcl_aux: Then, using the previous result, we prove sem_bullet by induction on i. Finally, sem_pcl_aux and sem_bullet give sem_pcl. *) -lemma LcatE : ∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a e2} a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{e1} a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/ae2|} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{e2}. +img class="anchor" src="icons/tick.png" id="LcatE"lemma LcatE : ∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a e2a title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{e1a title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/ae2a title="forget" href="cic:/fakeuri.def(1)"|/aa title="in_l" href="cic:/fakeuri.def(1)"}/a a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{e2a title="in_pl" href="cic:/fakeuri.def(1)"}/a. // qed. -lemma sem_pcl_aux : ∀S.∀e1:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{a title="eclose" href="cic:/fakeuri.def(1)"•/ai2} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i2} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/ai2|} → - a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1} a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/ai2|} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i2}. +img class="anchor" src="icons/tick.png" id="sem_pcl_aux"lemma sem_pcl_aux : ∀S.∀e1:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{a title="eclose" href="cic:/fakeuri.def(1)"•/ai2a title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i2a title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/ai2a title="forget" href="cic:/fakeuri.def(1)"|/aa title="in_l" href="cic:/fakeuri.def(1)"}/a → + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i2a title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1a title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/ai2a title="forget" href="cic:/fakeuri.def(1)"|/aa title="in_l" href="cic:/fakeuri.def(1)"}/a a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i2a title="in_pl" href="cic:/fakeuri.def(1)"}/a. #S * #i1 #b1 #i2 cases b1 [2:#th >a href="cic:/matita/tutorial/chapter8/pcl_false.def(5)"pcl_false/a >a href="cic:/matita/tutorial/chapter7/sem_pre_false.def(9)"sem_pre_false/a >a href="cic:/matita/tutorial/chapter7/sem_pre_false.def(9)"sem_pre_false/a >a href="cic:/matita/tutorial/chapter7/sem_cat.def(8)"sem_cat/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a/span/span/ |#H >a href="cic:/matita/tutorial/chapter8/pcl_true.def(5)"pcl_true/a >a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a @(a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a … (a href="cic:/matita/tutorial/chapter8/sem_pre_concat_r.def(10)"sem_pre_concat_r/a …)) @@ -190,8 +190,8 @@ lemma sem_pcl_aux : ∀S.∀e1:a href="cic:/matita/tutorial/chapter7/pre.def(1) ] qed. -lemma minus_eps_pre_aux: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀A. - a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i} a title="union" href="cic:/fakeuri.def(1)"∪/a A → a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i} a title="union" href="cic:/fakeuri.def(1)"∪/a (A a title="substraction" href="cic:/fakeuri.def(1)"-/a a title="singleton" href="cic:/fakeuri.def(1)"{/aa title="nil" href="cic:/fakeuri.def(1)"[/a ]}). +img class="anchor" src="icons/tick.png" id="minus_eps_pre_aux"lemma minus_eps_pre_aux: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀A. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{ea title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{ia title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="union" href="cic:/fakeuri.def(1)"∪/a A → a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a ea title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{ia title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="union" href="cic:/fakeuri.def(1)"∪/a (A a title="substraction" href="cic:/fakeuri.def(1)"-/a a title="singleton" href="cic:/fakeuri.def(1)"{/aa title="nil" href="cic:/fakeuri.def(1)"[/a a title="nil" href="cic:/fakeuri.def(1)"]/aa title="singleton" href="cic:/fakeuri.def(1)"}/a). #S #e #i #A #seme @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter7/minus_eps_pre.def(10)"minus_eps_pre/a] @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a [|@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter7/minus_eps_item.def(9)"minus_eps_item/a]] @@ -199,7 +199,7 @@ lemma minus_eps_pre_aux: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.de @a href="cic:/matita/tutorial/chapter4/eqP_substract_r.def(3)"eqP_substract_r/a // qed. -theorem sem_bull: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a. ∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{a title="eclose" href="cic:/fakeuri.def(1)"•/ai} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/ai|}. +img class="anchor" src="icons/tick.png" id="sem_bull"theorem sem_bull: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a. ∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{a title="eclose" href="cic:/fakeuri.def(1)"•/aia title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{ia title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/aia title="forget" href="cic:/fakeuri.def(1)"|/aa title="in_l" href="cic:/fakeuri.def(1)"}/a. #S #e elim e [#w normalize % [/span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/ | * //] |/span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a/span/span/ @@ -238,20 +238,20 @@ having e as carrier and no point, and then broadcast a point in it. The semantic (blank e) is obviously the empty language: from the point of view of the automaton, it corresponds with the pit state. *) -let rec blank (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (i: a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S) on i :a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S ≝ +img class="anchor" src="icons/tick.png" id="blank"let rec blank (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (i: a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S) on i :a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S ≝ match i with [ z ⇒ a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"pz/a ? | e ⇒ a title="pitem epsilon" href="cic:/fakeuri.def(1)"ϵ/a | s y ⇒ a title="pitem ps" href="cic:/fakeuri.def(1)"`/ay | o e1 e2 ⇒ (blank S e1) a title="pitem or" href="cic:/fakeuri.def(1)"+/a (blank S e2) | c e1 e2 ⇒ (blank S e1) a title="pitem cat" href="cic:/fakeuri.def(1)"·/a (blank S e2) - | k e ⇒ (blank S e)a title="pitem star" href="cic:/fakeuri.def(1)"^/a* ]. + | k e ⇒ (blank S e)a title="pitem star" href="cic:/fakeuri.def(1)"^/aa title="pitem star" href="cic:/fakeuri.def(1)"*/a ]. -lemma forget_blank: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S.a title="forget" href="cic:/fakeuri.def(1)"|/aa href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S e| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a e. +img class="anchor" src="icons/tick.png" id="forget_blank"lemma forget_blank: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S.a title="forget" href="cic:/fakeuri.def(1)"|/aa href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S ea title="forget" href="cic:/fakeuri.def(1)"|/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a e. #S #e elim e normalize // qed. -lemma sem_blank: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S.a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S e} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="empty set" href="cic:/fakeuri.def(1)"∅/a. +img class="anchor" src="icons/tick.png" id="sem_blank"lemma sem_blank: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S.a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S ea title="in_pl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="empty set" href="cic:/fakeuri.def(1)"∅/a. #S #e elim e [1,2:@a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a // |#s @a href="cic:/matita/tutorial/chapter8/eq_to_ex_eq.def(4)"eq_to_ex_eq/a // @@ -267,8 +267,8 @@ lemma sem_blank: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)" ] qed. -theorem re_embedding: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S. - a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{a title="eclose" href="cic:/fakeuri.def(1)"•/a(a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S e)} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{e}. +img class="anchor" src="icons/tick.png" id="re_embedding"theorem re_embedding: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{a title="eclose" href="cic:/fakeuri.def(1)"•/a(a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S e)a title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{ea title="in_l" href="cic:/fakeuri.def(1)"}/a. #S #e @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter8/sem_bull.def(12)"sem_bull/a] >a href="cic:/matita/tutorial/chapter8/forget_blank.def(4)"forget_blank/a @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a [|@a href="cic:/matita/tutorial/chapter8/sem_blank.def(9)"sem_blank/a]] @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/union_comm.def(3)"union_comm/a] @a href="cic:/matita/tutorial/chapter4/union_empty_r.def(3)"union_empty_r/a. @@ -280,36 +280,36 @@ qed. Plus and bullet have been already lifted from items to pres. We can now do a similar job for concatenation ⊙ and Kleene's star ⊛.*) -definition lifted_cat ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. +img class="anchor" src="icons/tick.png" id="lifted_cat"definition lifted_cat ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a href="cic:/matita/tutorial/chapter8/lift.def(2)"lift/a S (a href="cic:/matita/tutorial/chapter8/pre_concat_l.def(3)"pre_concat_l/a S a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"eclose/a e). notation "e1 ⊙ e2" left associative with precedence 70 for @{'odot $e1 $e2}. interpretation "lifted cat" 'odot e1 e2 = (lifted_cat ? e1 e2). -lemma odot_true_b : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀b. - a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai2,b〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2)),a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a b〉. +img class="anchor" src="icons/tick.png" id="odot_true_b"lemma odot_true_b : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀b. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai2,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2)),a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. #S #i1 #i2 #b normalize in ⊢ (??%?); cases (a title="eclose" href="cic:/fakeuri.def(1)"•/ai2) // qed. -lemma odot_false_b : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀b. - a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai2,b〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2 ,b〉. +img class="anchor" src="icons/tick.png" id="odot_false_b"lemma odot_false_b : ∀S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀b. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai2,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2 ,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. // qed. -lemma erase_odot:∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a e2)| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1| a title="re cat" href="cic:/fakeuri.def(1)"·/a (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2|). +img class="anchor" src="icons/tick.png" id="erase_odot"lemma erase_odot:∀S.∀e1,e2:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a e2)a title="forget" href="cic:/fakeuri.def(1)"|/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e1a title="forget" href="cic:/fakeuri.def(1)"|/a a title="re cat" href="cic:/fakeuri.def(1)"·/a (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2a title="forget" href="cic:/fakeuri.def(1)"|/a). #S * #i1 * * #i2 #b2 // >a href="cic:/matita/tutorial/chapter8/odot_true_b.def(6)"odot_true_b/a >a href="cic:/matita/tutorial/chapter7/erase_dot.def(4)"erase_dot/a // qed. (* Let us come to the star operation: *) -definition lk ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. +img class="anchor" src="icons/tick.png" id="lk"definition lk ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. match e with [ mk_Prod i1 b1 ⇒ match b1 with - [true ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"eclose/a ? i1))a title="pitem star" href="cic:/fakeuri.def(1)"^/a*, a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉 - |false ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1a title="pitem star" href="cic:/fakeuri.def(1)"^/a*,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 + [true ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"eclose/a ? i1))a title="pitem star" href="cic:/fakeuri.def(1)"^/aa title="pitem star" href="cic:/fakeuri.def(1)"*/a, a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a + |false ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai1a title="pitem star" href="cic:/fakeuri.def(1)"^/aa title="pitem star" href="cic:/fakeuri.def(1)"*/a,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a ] ]. @@ -317,22 +317,22 @@ definition lk ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)" interpretation "lk" 'lk a = (lk ? a). notation "a^⊛" non associative with precedence 90 for @{'lk $a}. -lemma ostar_true: ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉a title="lk" href="cic:/fakeuri.def(1)"^/a⊛ a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai))a title="pitem star" href="cic:/fakeuri.def(1)"^/a*, a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉. +img class="anchor" src="icons/tick.png" id="ostar_true"lemma ostar_true: ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/aa title="lk" href="cic:/fakeuri.def(1)"^/aa title="lk" href="cic:/fakeuri.def(1)"⊛/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a(a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a title="eclose" href="cic:/fakeuri.def(1)"•/ai))a title="pitem star" href="cic:/fakeuri.def(1)"^/aa title="pitem star" href="cic:/fakeuri.def(1)"*/a, a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. // qed. -lemma ostar_false: ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉a title="lk" href="cic:/fakeuri.def(1)"^/a⊛ a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aia title="pitem star" href="cic:/fakeuri.def(1)"^/a*, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉. +img class="anchor" src="icons/tick.png" id="ostar_false"lemma ostar_false: ∀S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/aa title="lk" href="cic:/fakeuri.def(1)"^/aa title="lk" href="cic:/fakeuri.def(1)"⊛/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aia title="pitem star" href="cic:/fakeuri.def(1)"^/aa title="pitem star" href="cic:/fakeuri.def(1)"*/a, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. // qed. -lemma erase_ostar: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (ea title="lk" href="cic:/fakeuri.def(1)"^/a⊛)| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|a title="re star" href="cic:/fakeuri.def(1)"^/a*. +img class="anchor" src="icons/tick.png" id="erase_ostar"lemma erase_ostar: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (ea title="lk" href="cic:/fakeuri.def(1)"^/aa title="lk" href="cic:/fakeuri.def(1)"⊛/a)a title="forget" href="cic:/fakeuri.def(1)"|/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a ea title="forget" href="cic:/fakeuri.def(1)"|/aa title="re star" href="cic:/fakeuri.def(1)"^/aa title="re star" href="cic:/fakeuri.def(1)"*/a. #S * #i * // qed. -lemma sem_odot_true: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i. - a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="singleton" href="cic:/fakeuri.def(1)"{/a a title="nil" href="cic:/fakeuri.def(1)"[/a ] }. +img class="anchor" src="icons/tick.png" id="sem_odot_true"lemma sem_odot_true: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/aa title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a ia title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="union" href="cic:/fakeuri.def(1)"∪/a a title="singleton" href="cic:/fakeuri.def(1)"{/a a title="nil" href="cic:/fakeuri.def(1)"[/a a title="nil" href="cic:/fakeuri.def(1)"]/a a title="singleton" href="cic:/fakeuri.def(1)"}/a. #S #e1 #i -cut (e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a(e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a〉) [//] +cut (e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a(e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) [//] #H >H cases (e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i) #i1 #b1 cases b1 [>a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [||@a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a @a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"union_assoc/a] @a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"eqP_union_l/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"eqP_sym/a/span/span/ @@ -340,18 +340,18 @@ cut (e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair ] qed. -lemma eq_odot_false: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i. - e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i. +img class="anchor" src="icons/tick.png" id="eq_odot_false"lemma eq_odot_false: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e1:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S.∀i. + e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i. #S #e1 #i -cut (e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a(e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉) [//] +cut (e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i), a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a(e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i) a title="boolean or" href="cic:/fakeuri.def(1)"∨/a a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a) [//] cases (e1 a title="item-pre concat" href="cic:/fakeuri.def(1)"▹/a i) #i1 #b1 cases b1 #H @H qed. (* We conclude this section with the proof of the main semantic properties of ⊙ and ⊛. *) -lemma sem_odot: - ∀S.∀e1,e2: a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a e2} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1}a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2|} a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2}. +img class="anchor" src="icons/tick.png" id="sem_odot"lemma sem_odot: + ∀S.∀e1,e2: a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1 a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a e2a title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e1a title="in_prl" href="cic:/fakeuri.def(1)"}/aa title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e2a title="forget" href="cic:/fakeuri.def(1)"|/aa title="in_l" href="cic:/fakeuri.def(1)"}/a a title="union" href="cic:/fakeuri.def(1)"∪/a a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e2a title="in_prl" href="cic:/fakeuri.def(1)"}/a. #S #e1 * #i2 * [>a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter8/sem_odot_true.def(10)"sem_odot_true/a] @@ -360,8 +360,8 @@ lemma sem_odot: ] qed. -theorem sem_ostar: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{ea title="lk" href="cic:/fakeuri.def(1)"^/a⊛} a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e} a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|}a title="star lang" href="cic:/fakeuri.def(1)"^/a*. +img class="anchor" src="icons/tick.png" id="sem_ostar"theorem sem_ostar: ∀S.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{ea title="lk" href="cic:/fakeuri.def(1)"^/aa title="lk" href="cic:/fakeuri.def(1)"⊛/aa title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="extensional equality" href="cic:/fakeuri.def(1)"=/a1 a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{ea title="in_prl" href="cic:/fakeuri.def(1)"}/a a title="cat lang" href="cic:/fakeuri.def(1)"·/a a title="in_l" href="cic:/fakeuri.def(1)"\sem/a{a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a ea title="forget" href="cic:/fakeuri.def(1)"|/aa title="in_l" href="cic:/fakeuri.def(1)"}/aa title="star lang" href="cic:/fakeuri.def(1)"^/aa title="star lang" href="cic:/fakeuri.def(1)"*/a. #S * #i #b cases b [>a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a >a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"sem_pre_true/a >a href="cic:/matita/tutorial/chapter7/sem_star.def(8)"sem_star/a >a href="cic:/matita/tutorial/chapter8/erase_bull.def(6)"erase_bull/a @a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"eqP_trans/a [|@a href="cic:/matita/tutorial/chapter4/eqP_union_r.def(3)"eqP_union_r/a[|@a href="cic:/matita/tutorial/chapter6/cat_ext_l.def(5)"cat_ext_l/a [|@a href="cic:/matita/tutorial/chapter8/minus_eps_pre_aux.def(11)"minus_eps_pre_aux/a //]]] diff --git a/weblib/tutorial/chapter9.ma b/weblib/tutorial/chapter9.ma index 13a0526b1..51fe4a402 100644 --- a/weblib/tutorial/chapter9.ma +++ b/weblib/tutorial/chapter9.ma @@ -11,26 +11,26 @@ lifted operators of the previous section: include "tutorial/chapter8.ma". -let rec move (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (x:S) (E: a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S) on E : a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S ≝ +img class="anchor" src="icons/tick.png" id="move"let rec move (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (x:S) (E: a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S) on E : a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S ≝ match E with - [ pz ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"pz/a S, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a 〉 - | pe ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem epsilon" href="cic:/fakeuri.def(1)"ϵ/a, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a 〉 - | ps y ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem ps" href="cic:/fakeuri.def(1)"`/ay, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a 〉 - | pp y ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem ps" href="cic:/fakeuri.def(1)"`/ay, x a title="eqb" href="cic:/fakeuri.def(1)"=/a= y 〉 + [ pz ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"pz/a S, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a + | pe ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem epsilon" href="cic:/fakeuri.def(1)"ϵ/a, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a + | ps y ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem ps" href="cic:/fakeuri.def(1)"`/ay, a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a + | pp y ⇒ a title="Pair construction" href="cic:/fakeuri.def(1)"〈/a a title="pitem ps" href="cic:/fakeuri.def(1)"`/ay, x a title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/a y a title="Pair construction" href="cic:/fakeuri.def(1)"〉/a | po e1 e2 ⇒ (move ? x e1) a title="oplus" href="cic:/fakeuri.def(1)"⊕/a (move ? x e2) | pc e1 e2 ⇒ (move ? x e1) a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a (move ? x e2) - | pk e ⇒ (move ? x e)a title="lk" href="cic:/fakeuri.def(1)"^/a⊛ ]. + | pk e ⇒ (move ? x e)a title="lk" href="cic:/fakeuri.def(1)"^/aa title="lk" href="cic:/fakeuri.def(1)"⊛/a ]. -lemma move_plus: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀x:S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. +img class="anchor" src="icons/tick.png" id="move_plus"lemma move_plus: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀x:S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S x (i1 a title="pitem or" href="cic:/fakeuri.def(1)"+/a i2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? x i1) a title="oplus" href="cic:/fakeuri.def(1)"⊕/a (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? x i2). // qed. -lemma move_cat: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀x:S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. +img class="anchor" src="icons/tick.png" id="move_cat"lemma move_cat: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀x:S.∀i1,i2:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S x (i1 a title="pitem cat" href="cic:/fakeuri.def(1)"·/a i2) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? x i1) a title="lifted cat" href="cic:/fakeuri.def(1)"⊙/a (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? x i2). // qed. -lemma move_star: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀x:S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S x ia title="pitem star" href="cic:/fakeuri.def(1)"^/a* a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? x i)a title="lk" href="cic:/fakeuri.def(1)"^/a⊛. +img class="anchor" src="icons/tick.png" id="move_star"lemma move_star: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀x:S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S x ia title="pitem star" href="cic:/fakeuri.def(1)"^/aa title="pitem star" href="cic:/fakeuri.def(1)"*/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? x i)a title="lk" href="cic:/fakeuri.def(1)"^/aa title="lk" href="cic:/fakeuri.def(1)"⊛/a. // qed. (* @@ -43,31 +43,31 @@ For a, we have two possible positions (all other points gets erased); the innerm point stops in front of the final b, while the other one broadcast inside (b^*a + b)b, so - move((•a + ϵ)((•b)*•a + •b)b,a) = 〈(a + ϵ)((•b)^*•a + •b)•b, false〉 + move((•a + ϵ)((•b)*•a + •b)b,a) = 〈(a + ϵ)((•b)^*•a + •b)•b, false〉 For b, we have two positions too. The innermost point stops in front of the final b too, while the other point reaches the end of b* and must go back through b*a: - move((•a + ϵ)((•b)*•a + •b)b ,b) = 〈(a + ϵ)((•b)*•a + b)•b, false〉 + move((•a + ϵ)((•b)*•a + •b)b ,b) = 〈(a + ϵ)((•b)*•a + b)•b, false〉 *) -definition pmove ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λx:S.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? x (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e). +img class="anchor" src="icons/tick.png" id="pmove"definition pmove ≝ λS:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.λx:S.λe:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? x (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e). -lemma pmove_def : ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀x:S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀b. - a href="cic:/matita/tutorial/chapter9/pmove.def(7)"pmove/a ? x a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,b〉 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? x i. +img class="anchor" src="icons/tick.png" id="pmove_def"lemma pmove_def : ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀x:S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀b. + a href="cic:/matita/tutorial/chapter9/pmove.def(7)"pmove/a ? x a title="Pair construction" href="cic:/fakeuri.def(1)"〈/ai,ba title="Pair construction" href="cic:/fakeuri.def(1)"〉/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? x i. // qed. -lemma eq_to_eq_hd: ∀A.∀l1,l2:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.∀a,b. - aa title="cons" href="cic:/fakeuri.def(1)":/a:l1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a ba title="cons" href="cic:/fakeuri.def(1)":/a:l2 → a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b. +img class="anchor" src="icons/tick.png" id="eq_to_eq_hd"lemma eq_to_eq_hd: ∀A.∀l1,l2:a href="cic:/matita/basics/list/list.ind(1,0,1)"list/a A.∀a,b. + aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al1 a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a ba title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/al2 → a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a b. #A #l1 #l2 #a #b #H destruct // qed. (* Obviously, a move does not change the carrier of the item, as one can easily prove by induction on the item. *) -lemma same_kernel: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a:S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. - a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? a i)| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/ai|. +img class="anchor" src="icons/tick.png" id="same_kernel"lemma same_kernel: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a:S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. + a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? a i)a title="forget" href="cic:/fakeuri.def(1)"|/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aia title="forget" href="cic:/fakeuri.def(1)"|/a. #S #a #i elim i // [#i1 #i2 #H1 #H2 >a href="cic:/matita/tutorial/chapter9/move_cat.def(7)"move_cat/a >a href="cic:/matita/tutorial/chapter8/erase_odot.def(7)"erase_odot/a // |#i1 #i2 #H1 #H2 >a href="cic:/matita/tutorial/chapter9/move_plus.def(7)"move_plus/a whd in ⊢ (??%%); // @@ -77,14 +77,14 @@ qed. (* Here is our first, major result, stating the correctness of the move operation. The proof is a simple induction on i. *) -theorem move_ok: +img class="anchor" src="icons/tick.png" id="move_ok"theorem move_ok: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a:S.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S.∀w: a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S. - a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? a i} w a title="iff" href="cic:/fakeuri.def(1)"↔/a a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{i} (aa title="cons" href="cic:/fakeuri.def(1)":/a:w). + a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a ? a ia title="in_prl" href="cic:/fakeuri.def(1)"}/a w a title="iff" href="cic:/fakeuri.def(1)"↔/a a title="in_pl" href="cic:/fakeuri.def(1)"\sem/a{ia title="in_pl" href="cic:/fakeuri.def(1)"}/a (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aw). #S #a #i elim i [normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ |normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ |normalize /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/And.con(0,1,2)"conj/a/span/span/ - |normalize #x #w cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x)) #H >H normalize + |normalize #x #w cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ax)) #H >H normalize [>(\P H) % [* // #bot @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a //| #H1 destruct /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,2,2)"or_intror/a/span/span/] |% [@a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a |#H1 cases (\Pf H) #H2 @H2 destruct //] ] @@ -107,37 +107,37 @@ qed. notation > "x ↦* E" non associative with precedence 60 for @{moves ? $x $E}. -let rec moves (S : a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) w e on w : a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S ≝ +img class="anchor" src="icons/tick.png" id="moves"let rec moves (S : a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) w e on w : a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S ≝ match w with [ nil ⇒ e | cons x w' ⇒ w' ↦* (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S x (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e))]. -lemma moves_empty: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? a title="nil" href="cic:/fakeuri.def(1)"[/a ] e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a e. +img class="anchor" src="icons/tick.png" id="moves_empty"lemma moves_empty: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? a title="nil" href="cic:/fakeuri.def(1)"[/a a title="nil" href="cic:/fakeuri.def(1)"]/a e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a e. // qed. -lemma moves_cons: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a:S.∀w.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? (aa title="cons" href="cic:/fakeuri.def(1)":/a:w) e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e)). +img class="anchor" src="icons/tick.png" id="moves_cons"lemma moves_cons: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a:S.∀w.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? (aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aw) e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w (a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e)). // qed. -lemma moves_left : ∀S,a,w,e. - a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S (wa title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a])) e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e)). +img class="anchor" src="icons/tick.png" id="moves_left"lemma moves_left : ∀S,a,w,e. + a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S (wa title="append" href="cic:/fakeuri.def(1)"@/a(aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a)) e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e)). #S #a #w elim w // #x #tl #Hind #e >a href="cic:/matita/tutorial/chapter9/moves_cons.def(8)"moves_cons/a >a href="cic:/matita/tutorial/chapter9/moves_cons.def(8)"moves_cons/a // qed. -lemma not_epsilon_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a:S.∀w: a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S. ∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a href="cic:/matita/basics/logic/iff.def(1)"iff/a ((aa title="cons" href="cic:/fakeuri.def(1)":/a:w) a title="in_prl mem" href="cic:/fakeuri.def(1)"∈/a e) ((aa title="cons" href="cic:/fakeuri.def(1)":/a:w) a title="in_pl mem" href="cic:/fakeuri.def(1)"∈/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e). +img class="anchor" src="icons/tick.png" id="not_epsilon_sem"lemma not_epsilon_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀a:S.∀w: a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S. ∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a href="cic:/matita/basics/logic/iff.def(1)"iff/a ((aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aw) a title="in_prl mem" href="cic:/fakeuri.def(1)"∈/a e) ((aa title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aw) a title="in_pl mem" href="cic:/fakeuri.def(1)"∈/a a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e). #S #a #w * #i #b cases b normalize [% /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/Or.con(0,1,2)"or_introl/a/span/span/ * // #H destruct |% normalize /span class="autotactic"2span class="autotrace" trace /span/span/] qed. -lemma same_kernel_moves: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀w.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e)| a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|. +img class="anchor" src="icons/tick.png" id="same_kernel_moves"lemma same_kernel_moves: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀w.∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e)a title="forget" href="cic:/fakeuri.def(1)"|/a a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a ea title="forget" href="cic:/fakeuri.def(1)"|/a. #S #w elim w // qed. -theorem decidable_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀w: a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S. ∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. - (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) a title="iff" href="cic:/fakeuri.def(1)"↔/a a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{e} w. +img class="anchor" src="icons/tick.png" id="decidable_sem"theorem decidable_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a.∀w: a href="cic:/matita/tutorial/chapter6/word.def(3)"word/a S. ∀e:a href="cic:/matita/tutorial/chapter7/pre.def(1)"pre/a S. + (a title="pair pi2" href="cic:/fakeuri.def(1)"\snd/a (a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a ? w e) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a) a title="iff" href="cic:/fakeuri.def(1)"↔/a a title="in_prl" href="cic:/fakeuri.def(1)"\sem/a{ea title="in_prl" href="cic:/fakeuri.def(1)"}/a w. #S #w elim w [* #i #b >a href="cic:/matita/tutorial/chapter9/moves_empty.def(8)"moves_empty/a cases b % /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/tutorial/chapter7/true_to_epsilon.def(9)"true_to_epsilon/a/span/span/span class="error" title="error location"/span #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ |#a #w1 #Hind #e >a href="cic:/matita/tutorial/chapter9/moves_cons.def(8)"moves_cons/a @@ -147,7 +147,7 @@ theorem decidable_sem: ∀S:a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1, qed. (* It is now clear that we can build a DFA D_e for e by taking pre as states, -and move as transition function; the initial state is •(e) and a state 〈i,b〉 is +and move as transition function; the initial state is •(e) and a state 〈i,b〉 is final if and only if b is true. The fact that states in D_e are finite is obvious: in fact, their cardinality is at most 2^{n+1} where n is the number of symbols in e. This is one of the advantages of pointed regular expressions w.r.t. derivatives, @@ -162,7 +162,7 @@ Below is the DFA associated with the regular expression (ac+bc)*. The graphical description of the automaton is the traditional one, with nodes for states and labelled arcs for transitions. Unreachable states are not shown. -Final states are emphasized by a double circle: since a state 〈e,b〉 is final if and +Final states are emphasized by a double circle: since a state 〈e,b〉 is final if and only if b is true, we may just label nodes with the item. The automaton is not minimal: it is easy to see that the two states corresponding to the items (a•c +bc)* and (ac+b•c)* are equivalent (a way to prove it is to observe @@ -188,16 +188,16 @@ are final, while 8 and 9 are not). We conclude this chapter with a few properties of the move opertions in relation with the pit state. *) -definition pit_pre ≝ λS.λi.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S (a title="forget" href="cic:/fakeuri.def(1)"|/ai|), a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/a〉. +img class="anchor" src="icons/tick.png" id="pit_pre"definition pit_pre ≝ λS.λi.a title="Pair construction" href="cic:/fakeuri.def(1)"〈/aa href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"blank/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aia title="forget" href="cic:/fakeuri.def(1)"|/a), a href="cic:/matita/basics/bool/bool.con(0,2,0)"false/aa title="Pair construction" href="cic:/fakeuri.def(1)"〉/a. (* The following function compute the list of characters occurring in a given item i. *) -let rec occur (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (i: a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S) on i ≝ +img class="anchor" src="icons/tick.png" id="occur"let rec occur (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqSet/a) (i: a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"re/a S) on i ≝ match i with - [ z ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/a ] - | e ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/a ] - | s y ⇒ ya title="cons" href="cic:/fakeuri.def(1)":/a:a title="nil" href="cic:/fakeuri.def(1)"[/a] + [ z ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/a a title="nil" href="cic:/fakeuri.def(1)"]/a + | e ⇒ a title="nil" href="cic:/fakeuri.def(1)"[/a a title="nil" href="cic:/fakeuri.def(1)"]/a + | s y ⇒ ya title="cons" href="cic:/fakeuri.def(1)":/aa title="cons" href="cic:/fakeuri.def(1)":/aa title="nil" href="cic:/fakeuri.def(1)"[/aa title="nil" href="cic:/fakeuri.def(1)"]/a | o e1 e2 ⇒ a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (occur S e1) (occur S e2) | c e1 e2 ⇒ a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"unique_append/a ? (occur S e1) (occur S e2) | k e ⇒ occur S e]. @@ -205,10 +205,10 @@ let rec occur (S: a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"DeqS (* If a symbol a does not occur in i, then move(i,a) gets to the pit state. *) -lemma not_occur_to_pit: ∀S,a.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/ai|)) a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → +img class="anchor" src="icons/tick.png" id="not_occur_to_pit"lemma not_occur_to_pit: ∀S,a.∀i:a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"pitem/a S. a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aia title="forget" href="cic:/fakeuri.def(1)"|/a)) a title="leibnitz's non-equality" href="cic:/fakeuri.def(1)"≠/a a href="cic:/matita/basics/bool/bool.con(0,1,0)"true/a → a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a i a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i. #S #a #i elim i // - [#x normalize cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/a=x) normalize // #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ + [#x normalize cases (aa title="eqb" href="cic:/fakeuri.def(1)"=/aa title="eqb" href="cic:/fakeuri.def(1)"=/ax) normalize // #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ |#i1 #i2 #Hind1 #Hind2 #H >a href="cic:/matita/tutorial/chapter9/move_cat.def(7)"move_cat/a >Hind1 [2:@(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 @a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l1.def(6)"sublist_unique_append_l1/a //] >Hind2 [2:@(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 @a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l2.def(6)"sublist_unique_append_l2/a //] // @@ -221,7 +221,7 @@ qed. (* We cannot escape form the pit state. *) -lemma move_pit: ∀S,a,i. a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i. +img class="anchor" src="icons/tick.png" id="move_pit"lemma move_pit: ∀S,a,i. a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6)"move/a S a (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a (a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i)) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i. #S #a #i elim i // [#i1 #i2 #Hind1 #Hind2 >a href="cic:/matita/tutorial/chapter9/move_cat.def(7)"move_cat/a >Hind1 >Hind2 // |#i1 #i2 #Hind1 #Hind2 >a href="cic:/matita/tutorial/chapter9/move_plus.def(7)"move_plus/a >Hind1 >Hind2 // @@ -229,18 +229,18 @@ lemma move_pit: ∀S,a,i. a href="cic:/matita/tutorial/chapter9/move.fix(0,2,6) ] qed. -lemma moves_pit: ∀S,w,i. a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w (a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i. +img class="anchor" src="icons/tick.png" id="moves_pit"lemma moves_pit: ∀S,w,i. a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w (a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i) a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S i. #S #w #i elim w // qed. (* If any character in w does not occur in i, then moves(i,w) gets to the pit state. *) -lemma to_pit: ∀S,w,e. a title="logical not" href="cic:/fakeuri.def(1)"¬/a a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S w (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|)) → +img class="anchor" src="icons/tick.png" id="to_pit"lemma to_pit: ∀S,w,e. a title="logical not" href="cic:/fakeuri.def(1)"¬/a a href="cic:/matita/tutorial/chapter5/sublist.def(5)"sublist/a S w (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a ea title="forget" href="cic:/fakeuri.def(1)"|/a)) → a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"moves/a S w e a title="leibnitz's equality" href="cic:/fakeuri.def(1)"=/a a href="cic:/matita/tutorial/chapter9/pit_pre.def(4)"pit_pre/a S (a title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e). #S #w elim w [#e * #H @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a @H normalize #a #abs @a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"False_ind/a /span class="autotactic"2span class="autotrace" trace a href="cic:/matita/basics/logic/absurd.def(2)"absurd/a/span/span/ - |#a #tl #Hind #e #H cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a e|)))) + |#a #tl #Hind #e #H cases (a href="cic:/matita/basics/bool/true_or_false.def(1)"true_or_false/a (a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"memb/a S a (a href="cic:/matita/tutorial/chapter9/occur.fix(0,1,6)"occur/a S (a title="forget" href="cic:/fakeuri.def(1)"|/aa title="pair pi1" href="cic:/fakeuri.def(1)"\fst/a ea title="forget" href="cic:/fakeuri.def(1)"|/a)))) [#Htrue >a href="cic:/matita/tutorial/chapter9/moves_cons.def(8)"moves_cons/a whd in ⊢ (???%); <(a href="cic:/matita/tutorial/chapter9/same_kernel.def(8)"same_kernel/a … a) @Hind >a href="cic:/matita/tutorial/chapter9/same_kernel.def(8)"same_kernel/a @(a href="cic:/matita/basics/logic/not_to_not.def(3)"not_to_not/a … H) #H1 #b #memb cases (a href="cic:/matita/basics/bool/orb_true_l.def(2)"orb_true_l/a … memb) [#H2 >(\P H2) // |#H2 @H1 //]