]> matita.cs.unibo.it Git - helm.git/commitdiff
commit by user utente2
authormatitaweb <claudio.sacerdoticoen@unibo.it>
Tue, 6 Mar 2012 18:01:38 +0000 (18:01 +0000)
committermatitaweb <claudio.sacerdoticoen@unibo.it>
Tue, 6 Mar 2012 18:01:38 +0000 (18:01 +0000)
weblib/tutorial/chapter7.ma
weblib/tutorial/chapter8.ma

index 17316f03169fc9a3707999b29b69298d9612ab8a..55a830179bdacb4e54c98f24e8a14ec67101554c 100644 (file)
@@ -229,11 +229,11 @@ interpretation "in_prl" 'in_l E = (in_prl ? E).
 (* The following, trivial lemmas are only meant for rewriting purposes. *)
 
 lemma sem_pre_true : ∀S.∀i:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S. 
-  \ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="in_pl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{i} \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="singleton" href="cic:/fakeuri.def(1)"\ 6{\ 5/a\ 6\ 5a title="epsilon" href="cic:/fakeuri.def(1)"\ 6ϵ\ 5/a\ 6}. 
+  \ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="in_pl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{i} \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="singleton" href="cic:/fakeuri.def(1)"\ 6{\ 5/a\ 6\ 5a title="epsilon" href="cic:/fakeuri.def(1)"\ 6ϵ\ 5/a\ 6}. 
 // qed.
 
 lemma sem_pre_false : ∀S.∀i:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S. 
-  \ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="in_pl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{i}. 
+  \ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="in_pl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{i}. 
 // qed.
 
 lemma sem_cat: ∀S.∀i1,i2:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S. 
index 7462d74b7717a254898c1c0fb9cf60673c19bcc7..a84d10a2a53223040e5d88c560203869bfc53a20 100644 (file)
@@ -13,7 +13,7 @@ reached the end of a + ϵ and we must pursue broadcasting inside (b*a + b)b. Aga
 parallel on the two additive subterms b^*a and b; the first point is allowed to both enter the 
 star, and to traverse it, stopping in front of a; the second point just stops in front of b. 
 No point reached that end of b^*a + b hence no further propagation is possible. In conclusion: 
-               •((a + ϵ)(b^*a + b)b) = 〈(•a + ϵ)((•b)^*•a + •b)b, false〉
+               •((a + ϵ)(b^*a + b)b) = 〈(•a + ϵ)((•b)^*•a + •b)b, false〉
 *)
 
 include "tutorial/chapter7.ma".
@@ -22,17 +22,17 @@ include "tutorial/chapter7.ma".
 the end of the expression. 
 Broadcasting inside a i1+i2 amounts to broadcast in parallel inside i1 and i2.
 If we define
-                 〈i1,b1〉 ⊕ 〈i2,b2〉 = 〈i1 + i2, b1∨ b2〉
+                 〈i1,b1〉 ⊕ 〈i2,b2〉 = 〈i1 + i2, b1∨ b2〉
 then, we just have •(i1+i2) = •(i1)⊕ •(i2).
 *)
 
 include "tutorial/chapter7.ma".
 
-definition lo ≝ λS:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.λa,b:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 a \ 5a title="pitem or" href="cic:/fakeuri.def(1)"\ 6+\ 5/a\ 6 \ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 b,\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 a \ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 b〉.
+definition lo ≝ λS:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.λa,b:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 a \ 5a title="pitem or" href="cic:/fakeuri.def(1)"\ 6+\ 5/a\ 6 \ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 b,\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 a \ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 b〉.
 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:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.∀b1,b2. \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,b1〉\ 5a title="oplus" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2,b2〉\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1\ 5a title="pitem or" href="cic:/fakeuri.def(1)"\ 6+\ 5/a\ 6i2,b1\ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6b2〉.
+lemma lo_def: ∀S.∀i1,i2:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.∀b1,b2. \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,b1〉\ 5a title="oplus" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2,b2〉\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1\ 5a title="pitem or" href="cic:/fakeuri.def(1)"\ 6+\ 5/a\ 6i2,b1\ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6b2〉.
 // qed.
 
 (*
@@ -42,15 +42,15 @@ point reached the end of i1. This suggests to define •(i1 · i2) as •(i1) 
 e ▹ i is a general operation of concatenation between a pre and an item, defined by 
 cases on the boolean in e: 
 
-       〈i1,true〉 ▹ i2  = i1 ◃ •(i_2)
-       〈i1,false〉 ▹ i2 = i1 · i2
+       〈i1,true〉 ▹ i2  = i1 ◃ •(i_2)
+       〈i1,false〉 ▹ i2 = i1 · i2
 In turn, ◃ says how to concatenate an item with a pre, that is however extremely simple:
-        i1 ◃ 〈i1,b〉  = 〈i_1 · i2, b〉
+        i1 ◃ 〈i1,b〉  = 〈i_1 · i2, b〉
 Let us come to the formalized definitions:
 *)
 
 definition pre_concat_r ≝ λS:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.λi:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.λe:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.
-  match e with [ mk_Prod i1 b ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 i1, b〉].
+  match e with [ mk_Prod i1 b ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 i1, b〉].
  
 notation "i ◃ e" left associative with precedence 60 for @{'lhd $i $e}.
 interpretation "pre_concat_r" 'lhd i e = (pre_concat_r ? i e).
@@ -77,7 +77,7 @@ definition pre_concat_l ≝ λS:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.in
   match e1 with 
   [ mk_Prod i1 b1 ⇒ match b1 with 
     [ true ⇒ (i1 \ 5a title="pre_concat_r" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (bcast ? i2)) 
-    | false ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1 \ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 i2,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6
+    | false ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1 \ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 i2,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6
     ]
   ].
 
@@ -90,13 +90,13 @@ notation "•" non associative with precedence 60 for @{eclose ?}.
 
 let rec eclose (S: \ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6) (i: \ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S) on i : \ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S ≝
  match i with
-  [ pz ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"\ 6pz\ 5/a\ 6 ?, \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6 〉
-  | pe ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="pitem epsilon" href="cic:/fakeuri.def(1)"\ 6ϵ\ 5/a\ 6,  \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 〉
-  | ps x ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="pitem pp" href="cic:/fakeuri.def(1)"\ 6`\ 5/a\ 6.x, \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6
-  | pp x ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="pitem pp" href="cic:/fakeuri.def(1)"\ 6`\ 5/a\ 6.x, \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6 〉
+  [ pz ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"\ 6pz\ 5/a\ 6 ?, \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6 〉
+  | pe ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="pitem epsilon" href="cic:/fakeuri.def(1)"\ 6ϵ\ 5/a\ 6,  \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 〉
+  | ps x ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="pitem pp" href="cic:/fakeuri.def(1)"\ 6`\ 5/a\ 6.x, \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6
+  | pp x ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="pitem pp" href="cic:/fakeuri.def(1)"\ 6`\ 5/a\ 6.x, \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6 〉
   | po i1 i2 ⇒ •i1 \ 5a title="oplus" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 •i2
   | pc i1 i2 ⇒ •i1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i2
-  | pk i ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (•i))\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6].
+  | pk i ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (•i))\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6].
   
 notation "• x" non associative with precedence 60 for @{'eclose $x}.
 interpretation "eclose" 'eclose x = (eclose ? x).
@@ -104,15 +104,15 @@ interpretation "eclose" 'eclose x = (eclose ? x).
 (* Here are a few simple properties of ▹ and •(-) *)
 
 lemma pcl_true : ∀S.∀i1,i2:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.
-  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 i1 \ 5a title="pre_concat_r" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2).
+  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 i1 \ 5a title="pre_concat_r" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2).
 // qed.
 
 lemma pcl_true_bis : ∀S.∀i1,i2:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.
-  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6〉 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1 \ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 \ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2), \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2)〉.
+  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6〉 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1 \ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 \ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2), \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2)〉.
 #S #i1 #i2 normalize cases (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2) // qed.
 
 lemma pcl_false: ∀S.∀i1,i2:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.
-  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6〉 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6  i2  \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1 \ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 i2, \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6.
+  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6〉 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6  i2  \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1 \ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 i2, \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6.
 // qed.
 
 lemma eclose_plus: ∀S:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.∀i1,i2:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.
@@ -124,7 +124,7 @@ lemma eclose_dot: ∀S:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"
 // qed.
 
 lemma eclose_star: ∀S:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.∀i:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.
-  \ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6(\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i))\ 5a title="pitem star" 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="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6\ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6(\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i))\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6.
 // qed.
 
 (* The definition of •(-) (eclose) can then be lifted from items to pres
@@ -132,7 +132,7 @@ in the obvious way. *)
 
 definition lift ≝ λS.λf:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S →\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.λe:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S. 
   match e with 
-  [ mk_Prod i b ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (f i), \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (f i) \ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 b〉].
+  [ mk_Prod i b ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (f i), \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (f i) \ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 b〉].
   
 definition preclose ≝ λS. \ 5a href="cic:/matita/tutorial/chapter8/lift.def(2)"\ 6lift\ 5/a\ 6 S (\ 5a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"\ 6eclose\ 5/a\ 6 S). 
 interpretation "preclose" 'eclose x = (preclose ? x).
@@ -290,12 +290,12 @@ 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:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.∀b. 
-  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6〉 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2,b〉 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1 \ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 (\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2)),\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2) \ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 b〉.
+  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6〉 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2,b〉 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1 \ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 (\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2)),\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2) \ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 b〉.
 #S #i1 #i2 #b normalize in ⊢ (??%?); cases (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2) // 
 qed.
 
 lemma odot_false_b : ∀S.∀i1,i2:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.∀b.
-  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6〉 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2,b〉 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1 \ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 i2 ,b〉.
+  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6〉 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i2,b〉 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1 \ 5a title="pitem cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 i2 ,b〉.
 // 
 qed.
   
@@ -310,8 +310,8 @@ definition lk ≝ λS:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6
   match e with 
   [ mk_Prod i1 b1 ⇒
     match b1 with 
-    [true ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"\ 6eclose\ 5/a\ 6 ? i1))\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*, \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6
-    |false ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6
+    [true ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter8/eclose.fix(0,1,4)"\ 6eclose\ 5/a\ 6 ? i1))\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*, \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6
+    |false ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i1\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6
     ]
   ]. 
 
@@ -320,11 +320,11 @@ interpretation "lk" 'lk a = (lk ? a).
 notation "a^⊛" non associative with precedence 90 for @{'lk $a}.
 
 lemma ostar_true: ∀S.∀i:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.
-  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6\ 5a title="lk" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6⊛ \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i))\ 5a title="pitem star" 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="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6\ 5a title="lk" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6⊛ \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i))\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*, \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6.
 // qed.
 
 lemma ostar_false: ∀S.∀i:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.
-  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6\ 5a title="lk" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6⊛ \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*, \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6.
+  \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6\ 5a title="lk" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6⊛ \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i\ 5a title="pitem star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*, \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6.
 // qed.
   
 lemma erase_ostar: ∀S.∀e:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.
@@ -332,9 +332,9 @@ lemma erase_ostar: ∀S.∀e:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6
 #S * #i * // qed.
 
 lemma sem_odot_true: ∀S:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.∀e1:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.∀i. 
-  \ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{e1 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6\ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6\ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i} \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="singleton" href="cic:/fakeuri.def(1)"\ 6{\ 5/a\ 6 \ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6 ] }.
+  \ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{e1 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6\ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6\ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i} \ 5a title="union" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="singleton" href="cic:/fakeuri.def(1)"\ 6{\ 5/a\ 6 \ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6 ] }.
 #S #e1 #i 
-cut (e1 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6〉 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i), \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6(e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i) \ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6) [//]
+cut (e1 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6〉 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i), \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6(e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i) \ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6) [//]
 #H >H cases (e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i) #i1 #b1 cases b1 
   [>\ 5a href="cic:/matita/tutorial/chapter7/sem_pre_true.def(9)"\ 6sem_pre_true\ 5/a\ 6 @\ 5a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"\ 6eqP_trans\ 5/a\ 6 [||@\ 5a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"\ 6eqP_sym\ 5/a\ 6 @\ 5a href="cic:/matita/tutorial/chapter4/union_assoc.def(3)"\ 6union_assoc\ 5/a\ 6]
    @\ 5a href="cic:/matita/tutorial/chapter4/eqP_union_l.def(3)"\ 6eqP_union_l\ 5/a\ 6 /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"\ 6eqP_sym\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6
@@ -343,9 +343,9 @@ cut (e1 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair
 qed.
 
 lemma eq_odot_false: ∀S:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.∀e1:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.∀i. 
-  e1 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i.
+  e1 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i.
 #S #e1 #i  
-cut (e1 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6〉 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i), \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6(e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i) \ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6) [//]
+cut (e1 \ 5a title="lifted cat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6〉 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i), \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6(e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i) \ 5a title="boolean or" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6) [//]
 cases (e1 \ 5a title="item-pre concat" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 i) #i1 #b1 cases b1 #H @H
 qed.