]> matita.cs.unibo.it Git - helm.git/commitdiff
commit by user andrea
authormatitaweb <claudio.sacerdoticoen@unibo.it>
Wed, 9 Nov 2011 10:13:47 +0000 (10:13 +0000)
committermatitaweb <claudio.sacerdoticoen@unibo.it>
Wed, 9 Nov 2011 10:13:47 +0000 (10:13 +0000)
weblib/tutorial/chapter4.ma

index 77e42066e2fc65801c5de91e505aac57ebda1dc5..00b70a0da09ab6eb691b838bb3733d0d161a1c69 100644 (file)
@@ -172,31 +172,18 @@ interpretation "in_prl" 'in_l E = (in_prl ? E).
 
 lemma not_epsilon_lp :∀S.∀pi:\ 5a href="cic:/matita/tutorial/chapter4/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S. \ 5a title="logical not" href="cic:/fakeuri.def(1)"\ 6¬\ 5/a\ 6 ((\ 5a title="in_pl" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 pi) \ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6]).
 #S #pi (elim pi) normalize /2/
-  [#pi1 #pi2 #H1 #H2 % * /2/ * #w1 * #w2 * * /2/;
-
-lemma append_eq_nil : ∀S.∀w1,w2:word S. w1 @ w2 = [ ] → w1 = [ ].
-#S w1; nelim w1; //. #x tl IH w2; nnormalize; #abs; ndestruct; nqed.
-
-(* lemma 12 *)
-nlemma epsilon_in_true : ∀S.∀e:pre S. [ ] ∈ e ↔ \snd e = true.
-#S r; ncases r; #e b; @; ##[##2: #H; nrewrite > H; @2; /2/; ##] ncases b;//; 
-nnormalize; *; ##[##2:*] nelim e;
-##[ ##1,2: *; ##| #c; *; ##| #c; nnormalize; #; ndestruct; ##| ##7: #p H;
-##| #r1 r2 H G; *; ##[##2: /3/ by or_intror]
-##| #r1 r2 H1 H2; *; /3/ by or_intror, or_introl; ##]
-*; #w1; *; #w2; *; *; #defw1; nrewrite > (append_eq_nil … w1 w2 …); /3/ by {};//;
-nqed.
-
-nlemma not_epsilon_lp : ∀S:Alpha.∀e:pitem S. ¬ ((𝐋\p e) [ ]).
-#S e; nelim e; nnormalize; /2/ by nmk;
-##[ #; @; #; ndestruct;
-##| #r1 r2 n1 n2; @; *; /2/; *; #w1; *; #w2; *; *; #H;
-    nrewrite > (append_eq_nil …H…); /2/;
-##| #r1 r2 n1 n2; @; *; /2/;
-##| #r n; @; *; #w1; *; #w2; *; *; #H;     
-    nrewrite > (append_eq_nil …H…); /2/;##]
-nqed.
-
-ndefinition lo ≝ λS:Alpha.λ
+  [#pi1 #pi2 #H1 #H2 % * /2/ * #w1 * #w2 * * #appnil 
+   cases (\ 5a href="cic:/matita/tutorial/chapter3/nil_to_nil.def(5)"\ 6nil_to_nil\ 5/a\ 6 … appnil) /2/
+  |#p11 #p12 #H1 #H2 % * /2/
+  |#pi #H % * #w1 * #w2 * * #appnil (cases (\ 5a href="cic:/matita/tutorial/chapter3/nil_to_nil.def(5)"\ 6nil_to_nil\ 5/a\ 6 … appnil)) /2/
+  ]
+qed.
+
+lemma if_true_epsilon: ∀S.∀e:\ 5a href="cic:/matita/tutorial/chapter4/pre.def(1)"\ 6pre\ 5/a\ 6 S. \ 5a title="snd" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 e \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 → ((\ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 e) \ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6]).
+#S #e #H %2 >H // qed.
+
+lemma if_epsilon_true : ∀S.∀e:\ 5a href="cic:/matita/tutorial/chapter4/pre.def(1)"\ 6pre\ 5/a\ 6 S. \ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6 ] \ 5a title="in_prl mem" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 e → \ 5a title="snd" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 e \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6.
+#S * #pi #b * [#abs @\ 5a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"\ 6False_ind\ 5/a\ 6 /2/] cases b normalize // @\ 5a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"\ 6False_ind\ 5/a\ 6
+qed.