]> matita.cs.unibo.it Git - helm.git/commitdiff
commit by user andrea
authormatitaweb <claudio.sacerdoticoen@unibo.it>
Wed, 7 Mar 2012 08:45:22 +0000 (08:45 +0000)
committermatitaweb <claudio.sacerdoticoen@unibo.it>
Wed, 7 Mar 2012 08:45:22 +0000 (08:45 +0000)
weblib/tutorial/chapter10.ma

index 9eef6a35099631f860949f2822eed3454847147a..24efb3e983e9f6986be23ed8055d09df08388b51 100644 (file)
@@ -79,11 +79,11 @@ lemma bisim_to_sem: ∀S:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0
   \ 5a href="cic:/matita/tutorial/chapter10/is_bisim.def(8)"\ 6is_bisim\ 5/a\ 6 S l (\ 5a href="cic:/matita/tutorial/chapter10/occ.def(7)"\ 6occ\ 5/a\ 6 S e1 e2) → \ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 ? \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6e1,e2〉 l \ 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\sem\ 5/a\ 6{e1}\ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 61\ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{e2}.
 #S #l #e1 #e2 #Hbisim #Hmemb @\ 5a href="cic:/matita/tutorial/chapter10/equiv_sem_occ.def(17)"\ 6equiv_sem_occ\ 5/a\ 6 
 #w #Hsub @(\ 5a href="cic:/matita/basics/logic/proj1.def(2)"\ 6proj1\ 5/a\ 6 … (Hbisim \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"\ 6moves\ 5/a\ 6 S w e1,\ 5a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"\ 6moves\ 5/a\ 6 S w e2〉 ?))
-lapply Hsub @(list_elim_left … w) [//]
-#a #w1 #Hind #Hsub >moves_left >moves_left @(proj2 …(Hbisim …(Hind ?)))
-  [#x #Hx @Hsub @memb_append_l1 //
-  |cut (memb S a (occ S e1 e2) = true) [@Hsub @memb_append_l2 //] #occa 
-   @(memb_map … occa)
+lapply Hsub @(\ 5a href="cic:/matita/basics/list/list_elim_left.def(10)"\ 6list_elim_left\ 5/a\ 6 … w) [//]
+#a #w1 #Hind #Hsub >\ 5a href="cic:/matita/tutorial/chapter9/moves_left.def(9)"\ 6moves_left\ 5/a\ 6 >\ 5a href="cic:/matita/tutorial/chapter9/moves_left.def(9)"\ 6moves_left\ 5/a\ 6 @(\ 5a href="cic:/matita/basics/logic/proj2.def(2)"\ 6proj2\ 5/a\ 6 …(Hbisim …(Hind ?)))
+  [#x #Hx @Hsub @\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l1.def(5)"\ 6memb_append_l1\ 5/a\ 6 //
+  |cut (\ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 S a (\ 5a href="cic:/matita/tutorial/chapter10/occ.def(7)"\ 6occ\ 5/a\ 6 S e1 e2) \ 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) [@Hsub @\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l2.def(5)"\ 6memb_append_l2\ 5/a\ 6 //] #occa 
+   @(\ 5a href="cic:/matita/tutorial/chapter5/memb_map.def(5)"\ 6memb_map\ 5/a\ 6 … occa)
   ]
 qed.
 
@@ -112,17 +112,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: list ?) on n ≝
+let rec bisim S l n (frontier,visited: \ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 ?) on n ≝
   match n with 
-  [ O ⇒ 〈false,visited〉 (* assert false *)
+  [ O ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6,visited〉 (* assert false *)
   | S m ⇒ 
     match frontier with
-    [ nil ⇒ 〈true,visited〉
+    [ nil ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6,visited〉
     | cons hd tl ⇒
-      if beqb (\snd (\fst hd)) (\snd (\snd hd)) then
-        bisim S l m (unique_append ? (filter ? (λx.notb (memb ? x (hd::visited))) 
-        (sons S l hd)) tl) (hd::visited)
-      else 〈false,visited〉
+      if \ 5a href="cic:/matita/tutorial/chapter4/beqb.def(2)"\ 6beqb\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 hd)) (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 hd)) then
+        bisim S l m (\ 5a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"\ 6unique_append\ 5/a\ 6 ? (\ 5a href="cic:/matita/basics/list/filter.def(2)"\ 6filter\ 5/a\ 6 ? (λx.\ 5a href="cic:/matita/basics/bool/notb.def(1)"\ 6notb\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 ? x (hd\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:visited))) 
+        (\ 5a href="cic:/matita/tutorial/chapter10/sons.def(7)"\ 6sons\ 5/a\ 6 S l hd)) tl) (hd\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:visited)
+      else \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6,visited〉
     ]
   ].
 
@@ -134,44 +134,44 @@ 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: list ?.
-  bisim S l n frontier visited =
+lemma unfold_bisim: ∀S,l,n.∀frontier,visited: \ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 ?.
+  \ 5a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"\ 6bisim\ 5/a\ 6 S l n frontier visited \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6
   match n with 
-  [ O ⇒ 〈false,visited〉 (* assert false *)
+  [ O ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6,visited〉 (* assert false *)
   | S m ⇒ 
     match frontier with
-    [ nil ⇒ 〈true,visited〉
+    [ nil ⇒ \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6,visited〉
     | cons hd tl ⇒
-      if beqb (\snd (\fst hd)) (\snd (\snd hd)) then
-        bisim S l m (unique_append ? (filter ? (λx.notb(memb ? x (hd::visited))) 
-          (sons S l hd)) tl) (hd::visited)
-      else 〈false,visited〉
+      if \ 5a href="cic:/matita/tutorial/chapter4/beqb.def(2)"\ 6beqb\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 hd)) (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 hd)) then
+        \ 5a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"\ 6bisim\ 5/a\ 6 S l m (\ 5a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"\ 6unique_append\ 5/a\ 6 ? (\ 5a href="cic:/matita/basics/list/filter.def(2)"\ 6filter\ 5/a\ 6 ? (λx.\ 5a href="cic:/matita/basics/bool/notb.def(1)"\ 6notb\ 5/a\ 6(\ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 ? x (hd\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:visited))) 
+          (\ 5a href="cic:/matita/tutorial/chapter10/sons.def(7)"\ 6sons\ 5/a\ 6 S l hd)) tl) (hd\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:visited)
+      else \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6\ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6,visited〉
     ]
   ].
 #S #l #n cases n // qed.
 
-lemma bisim_never: ∀S,l.∀frontier,visited: list ?.
-  bisim S l O frontier visited = 〈false,visited〉.
-#frontier #visited >unfold_bisim // 
+lemma bisim_never: ∀S,l.∀frontier,visited: \ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 ?.
+  \ 5a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"\ 6bisim\ 5/a\ 6 S l \ 5a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"\ 6O\ 5/a\ 6 frontier visited \ 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 href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6,visited〉.
+#frontier #visited >\ 5a href="cic:/matita/tutorial/chapter10/unfold_bisim.def(9)"\ 6unfold_bisim\ 5/a\ 6 // 
 qed.
 
-lemma bisim_end: ∀Sig,l,m.∀visited: list ?.
-  bisim Sig l (S m) [] visited = 〈true,visited〉.
-#n #visisted >unfold_bisim // 
+lemma bisim_end: ∀Sig,l,m.∀visited: \ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 ?.
+  \ 5a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"\ 6bisim\ 5/a\ 6 Sig l (\ 5a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"\ 6S\ 5/a\ 6 m) \ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6] visited \ 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 href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6,visited〉.
+#n #visisted >\ 5a href="cic:/matita/tutorial/chapter10/unfold_bisim.def(9)"\ 6unfold_bisim\ 5/a\ 6 // 
 qed.
 
-lemma bisim_step_true: ∀Sig,l,m.∀p.∀frontier,visited: list ?.
-beqb (\snd (\fst p)) (\snd (\snd p)) = true →
-  bisim Sig l (S m) (p::frontier) visited = 
-  bisim Sig l m (unique_append ? (filter ? (λx.notb(memb ? x (p::visited))) 
-    (sons Sig l p)) frontier) (p::visited).
-#Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // 
+lemma bisim_step_true: ∀Sig,l,m.∀p.∀frontier,visited: \ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 ?.
+\ 5a href="cic:/matita/tutorial/chapter4/beqb.def(2)"\ 6beqb\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 p)) (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 p)) \ 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 href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"\ 6bisim\ 5/a\ 6 Sig l (\ 5a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"\ 6S\ 5/a\ 6 m) (p\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:frontier) visited \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 
+  \ 5a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"\ 6bisim\ 5/a\ 6 Sig l m (\ 5a href="cic:/matita/tutorial/chapter5/unique_append.fix(0,1,5)"\ 6unique_append\ 5/a\ 6 ? (\ 5a href="cic:/matita/basics/list/filter.def(2)"\ 6filter\ 5/a\ 6 ? (λx.\ 5a href="cic:/matita/basics/bool/notb.def(1)"\ 6notb\ 5/a\ 6(\ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 ? x (p\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:visited))) 
+    (\ 5a href="cic:/matita/tutorial/chapter10/sons.def(7)"\ 6sons\ 5/a\ 6 Sig l p)) frontier) (p\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:visited).
+#Sig #l #m #p #frontier #visited #test >\ 5a href="cic:/matita/tutorial/chapter10/unfold_bisim.def(9)"\ 6unfold_bisim\ 5/a\ 6 whd in ⊢ (??%?);  >test // 
 qed.
 
-lemma bisim_step_false: ∀Sig,l,m.∀p.∀frontier,visited: list ?.
-beqb (\snd (\fst p)) (\snd (\snd p)) = false →
-  bisim Sig l (S m) (p::frontier) visited = 〈false,visited〉.
-#Sig #l #m #p #frontier #visited #test >unfold_bisim normalize nodelta >test // 
+lemma bisim_step_false: ∀Sig,l,m.∀p.∀frontier,visited: \ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 ?.
+\ 5a href="cic:/matita/tutorial/chapter4/beqb.def(2)"\ 6beqb\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 p)) (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 p)) \ 5a title="leibnitz's equality" 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 href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"\ 6bisim\ 5/a\ 6 Sig l (\ 5a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"\ 6S\ 5/a\ 6 m) (p\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:frontier) visited \ 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 href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6,visited〉.
+#Sig #l #m #p #frontier #visited #test >\ 5a href="cic:/matita/tutorial/chapter10/unfold_bisim.def(9)"\ 6unfold_bisim\ 5/a\ 6 whd in ⊢ (??%?); >test // 
 qed.
 
 (* In order to prove termination of bisim we must be able to effectively 
@@ -181,95 +181,95 @@ enumerate all possible pres: *)
 #b cases b normalize //
 qed. *)
 
-let rec pitem_enum S (i:re S) on i ≝
+let rec pitem_enum S (i:\ 5a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"\ 6re\ 5/a\ 6 S) on i ≝
   match i with
-  [ z ⇒ [pz S]
-  | e ⇒ [pe S]
-  | s y ⇒ [ps S y; pp S y]
-  | o i1 i2 ⇒ compose ??? (po S) (pitem_enum S i1) (pitem_enum S i2)
-  | c i1 i2 ⇒ compose ??? (pc S) (pitem_enum S i1) (pitem_enum S i2)
-  | k i ⇒ map ?? (pk S) (pitem_enum S i)
+  [ z ⇒ (\ 5a href="cic:/matita/tutorial/chapter7/pitem.con(0,1,1)"\ 6pz\ 5/a\ 6\ 5span class="error" title="Parse error: SYMBOL ':' or RPAREN expected after [term] (in [term])"\ 6\ 5/span\ 6\ 5span class="error" title="Parse error: SYMBOL ':' or RPAREN expected after [term] (in [term])"\ 6\ 5/span\ 6 S)\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:\ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6]
+  | e ⇒ (\ 5a href="cic:/matita/tutorial/chapter7/pitem.con(0,2,1)"\ 6pe\ 5/a\ 6 S)\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:\ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6]
+  | s y ⇒ (\ 5a href="cic:/matita/tutorial/chapter7/pitem.con(0,3,1)"\ 6ps\ 5/a\ 6 S y)\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:(\ 5a href="cic:/matita/tutorial/chapter7/pitem.con(0,4,1)"\ 6pp\ 5/a\ 6 S y)\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:\ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6]
+  | o i1 i2 ⇒ \ 5a href="cic:/matita/basics/list/compose.def(2)"\ 6compose\ 5/a\ 6 ??? (\ 5a href="cic:/matita/tutorial/chapter7/pitem.con(0,6,1)"\ 6po\ 5/a\ 6 S) (pitem_enum S i1) (pitem_enum S i2)
+  | c i1 i2 ⇒ \ 5a href="cic:/matita/basics/list/compose.def(2)"\ 6compose\ 5/a\ 6 ??? (\ 5a href="cic:/matita/tutorial/chapter7/pitem.con(0,5,1)"\ 6pc\ 5/a\ 6 S) (pitem_enum S i1) (pitem_enum S i2)
+  | k i ⇒ \ 5a href="cic:/matita/basics/list/map.fix(0,3,1)"\ 6map\ 5/a\ 6 ?? (\ 5a href="cic:/matita/tutorial/chapter7/pitem.con(0,7,1)"\ 6pk\ 5/a\ 6 S) (pitem_enum S i)
   ].
   
-lemma pitem_enum_complete : ∀S.∀i:pitem S.
-  memb (DeqItem S) i (pitem_enum S (|i|)) = true.
+lemma pitem_enum_complete : ∀S.∀i:\ 5a href="cic:/matita/tutorial/chapter7/pitem.ind(1,0,1)"\ 6pitem\ 5/a\ 6 S.
+  \ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter7/DeqItem.def(6)"\ 6DeqItem\ 5/a\ 6 S) i (\ 5a href="cic:/matita/tutorial/chapter10/pitem_enum.fix(0,1,3)"\ 6pitem_enum\ 5/a\ 6 S (\ 5a title="forget" href="cic:/fakeuri.def(1)"\ 6|\ 5/a\ 6i|)) \ 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 #i elim i 
   [1,2://
-  |3,4:#c normalize >(\b (refl … c)) //
-  |5,6:#i1 #i2 #Hind1 #Hind2 @(memb_compose (DeqItem S) (DeqItem S)) //
-  |#i #Hind @(memb_map (DeqItem S)) //
+  |3,4:#c normalize >(\b (\ 5a href="cic:/matita/basics/logic/eq.con(0,1,2)"\ 6refl\ 5/a\ 6 … c)) //
+  |5,6:#i1 #i2 #Hind1 #Hind2 @(\ 5a href="cic:/matita/tutorial/chapter5/memb_compose.def(6)"\ 6memb_compose\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter7/DeqItem.def(6)"\ 6DeqItem\ 5/a\ 6 S) (\ 5a href="cic:/matita/tutorial/chapter7/DeqItem.def(6)"\ 6DeqItem\ 5/a\ 6 S)) //
+  |#i #Hind @(\ 5a href="cic:/matita/tutorial/chapter5/memb_map.def(5)"\ 6memb_map\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter7/DeqItem.def(6)"\ 6DeqItem\ 5/a\ 6 S)) //
   ]
 qed.
 
-definition pre_enum ≝ λS.λi:re S.
-  compose ??? (λi,b.〈i,b〉) (pitem_enum S i) [true;false].
+definition pre_enum ≝ λS.λi:\ 5a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"\ 6re\ 5/a\ 6 S.
+  \ 5a href="cic:/matita/basics/list/compose.def(2)"\ 6compose\ 5/a\ 6 ??? (λi,b.\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,b〉) (\ 5a href="cic:/matita/tutorial/chapter10/pitem_enum.fix(0,1,3)"\ 6pitem_enum\ 5/a\ 6 S i) \ 5a title="cons" 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 href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6].
   
-lemma pre_enum_complete : ∀S.∀e:pre S.
-  memb ? e (pre_enum S (|\fst e|)) = true.
-#S * #i #b @(memb_compose (DeqItem S) DeqBool ? (λi,b.〈i,b〉))
+lemma pre_enum_complete : ∀S.∀e:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.
+  \ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 ? e (\ 5a href="cic:/matita/tutorial/chapter10/pre_enum.def(4)"\ 6pre_enum\ 5/a\ 6 S (\ 5a title="forget" href="cic:/fakeuri.def(1)"\ 6|\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 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 * #i #b @(\ 5a href="cic:/matita/tutorial/chapter5/memb_compose.def(6)"\ 6memb_compose\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter7/DeqItem.def(6)"\ 6DeqItem\ 5/a\ 6 S) \ 5a href="cic:/matita/tutorial/chapter4/DeqBool.def(5)"\ 6DeqBool\ 5/a\ 6 ? (λi,b.\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,b〉))
 // cases b normalize //
 qed.
  
-definition space_enum ≝ λS.λi1,i2:re S.
-  compose ??? (λe1,e2.〈e1,e2〉) (pre_enum S i1) (pre_enum S i2).
+definition space_enum ≝ λS.λi1,i2:\ 5a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"\ 6re\ 5/a\ 6 S.
+  \ 5a href="cic:/matita/basics/list/compose.def(2)"\ 6compose\ 5/a\ 6 ??? (λe1,e2.\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6e1,e2〉) (\ 5a href="cic:/matita/tutorial/chapter10/pre_enum.def(4)"\ 6pre_enum\ 5/a\ 6 S i1) (\ 5a href="cic:/matita/tutorial/chapter10/pre_enum.def(4)"\ 6pre_enum\ 5/a\ 6 S i2).
 
-lemma space_enum_complete : ∀S.∀e1,e2: pre S.
-  memb ? 〈e1,e2〉 (space_enum S (|\fst e1|) (|\fst e2|)) = true.
-#S #e1 #e2 @(memb_compose … (λi,b.〈i,b〉))
+lemma space_enum_complete : ∀S.∀e1,e2: \ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.
+  \ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 ? \ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6e1,e2〉 (\ 5a href="cic:/matita/tutorial/chapter10/space_enum.def(5)"\ 6space_enum\ 5/a\ 6 S (\ 5a title="forget" 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="forget" href="cic:/fakeuri.def(1)"\ 6|\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 e2|)) \ 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 #e1 #e2 @(\ 5a href="cic:/matita/tutorial/chapter5/memb_compose.def(6)"\ 6memb_compose\ 5/a\ 6 … (λi,b.\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6i,b〉))
 // qed.
 
-definition all_reachable ≝ λS.λe1,e2:pre S.λl: list ?.
-uniqueb ? l = true ∧ 
-  ∀p. memb ? p l = true → 
-    ∃w.(moves S w e1 = \fst p) ∧ (moves S w e2 = \snd p). 
+definition all_reachable ≝ λS.λe1,e2:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.λl: \ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 ?.
+\ 5a href="cic:/matita/tutorial/chapter5/uniqueb.fix(0,1,5)"\ 6uniqueb\ 5/a\ 6 ? l \ 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="logical and" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 
+  ∀p. \ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 ? p l \ 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="exists" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6w.(\ 5a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"\ 6moves\ 5/a\ 6 S w e1 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 p) \ 5a title="logical and" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"\ 6moves\ 5/a\ 6 S w e2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 p). 
 
-definition disjoint ≝ λS:DeqSet.λl1,l2.
-  ∀p:S. memb S p l1 = true →  memb S p l2 = false.
+definition disjoint ≝ λS:\ 5a href="cic:/matita/tutorial/chapter4/DeqSet.ind(1,0,0)"\ 6DeqSet\ 5/a\ 6.λl1,l2.
+  ∀p:S. \ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 S p l1 \ 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 href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 S p l2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6.
         
 (* We are ready to prove that bisim is correct; we use the invariant 
 that at each call of bisim the two lists visited and frontier only contain 
 nodes reachable from \langle e_1,e_2\rangle, hence it is absurd to suppose 
 to meet a pair which is not cofinal. *)
 
-lemma bisim_correct: ∀S.∀e1,e2:pre S.\sem{e1}=1\sem{e2} → 
- ∀l,n.∀frontier,visited:list ((pre S)×(pre S)).
|space_enum S (|\fst e1|) (|\fst e2|)| < n + |visited|→
all_reachable S e1 e2 visited →  
all_reachable S e1 e2 frontier →
disjoint ? frontier visited →
\fst (bisim S l n frontier visited) = true.
+lemma bisim_correct: ∀S.∀e1,e2:\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S.\ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{e1}\ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 61\ 5a title="in_prl" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{e2} → 
+ ∀l,n.∀frontier,visited:\ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 ((\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S)\ 5a title="Product" href="cic:/fakeuri.def(1)"\ 6×\ 5/a\ 6(\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S)).
\ 5a title="norm" href="cic:/fakeuri.def(1)"\ 6|\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/space_enum.def(5)"\ 6space_enum\ 5/a\ 6 S (\ 5a title="forget" 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="forget" href="cic:/fakeuri.def(1)"\ 6|\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 e2|)| \ 5a title="natural 'less than'" href="cic:/fakeuri.def(1)"\ 6<\ 5/a\ 6 n \ 5a title="natural plus" href="cic:/fakeuri.def(1)"\ 6+\ 5/a\ 6 \ 5a title="norm" href="cic:/fakeuri.def(1)"\ 6|\ 5/a\ 6visited|→
\ 5a href="cic:/matita/tutorial/chapter10/all_reachable.def(8)"\ 6all_reachable\ 5/a\ 6 S e1 e2 visited →  
\ 5a href="cic:/matita/tutorial/chapter10/all_reachable.def(8)"\ 6all_reachable\ 5/a\ 6 S e1 e2 frontier →
\ 5a href="cic:/matita/tutorial/chapter10/disjoint.def(5)"\ 6disjoint\ 5/a\ 6 ? frontier visited →
\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"\ 6bisim\ 5/a\ 6 S l n frontier visited) \ 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.
 #Sig #e1 #e2 #same #l #n elim n 
-  [#frontier #visited #abs * #unique #H @False_ind @(absurd … abs)
-   @le_to_not_lt @sublist_length // * #e11 #e21 #membp 
-   cut ((|\fst e11| = |\fst e1|) ∧ (|\fst e21| = |\fst e2|))
-   [|* #H1 #H2 <H1 <H2 @space_enum_complete]
-   cases (H … membp) #w * #we1 #we2 <we1 <we2 % >same_kernel_moves //    
-  |#m #HI * [#visited #vinv #finv >bisim_end //]
+  [#frontier #visited #abs * #unique #H @\ 5a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"\ 6False_ind\ 5/a\ 6 @(\ 5a href="cic:/matita/basics/logic/absurd.def(2)"\ 6absurd\ 5/a\ 6 … abs)
+   @\ 5a href="cic:/matita/arithmetics/nat/le_to_not_lt.def(8)"\ 6le_to_not_lt\ 5/a\ 6 @\ 5a href="cic:/matita/tutorial/chapter5/sublist_length.def(9)"\ 6sublist_length\ 5/a\ 6 // * #e11 #e21 #membp 
+   cut ((\ 5a title="forget" href="cic:/fakeuri.def(1)"\ 6|\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 e11| \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="forget" 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="logical and" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (\ 5a title="forget" href="cic:/fakeuri.def(1)"\ 6|\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 e21| \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="forget" href="cic:/fakeuri.def(1)"\ 6|\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 e2|))
+   [|* #H1 #H2 <H1 <H2 @\ 5a href="cic:/matita/tutorial/chapter10/space_enum_complete.def(9)"\ 6space_enum_complete\ 5/a\ 6]
+   cases (H … membp) #w * #we1 #we2 <we1 <we2 % >\ 5a href="cic:/matita/tutorial/chapter9/same_kernel_moves.def(9)"\ 6same_kernel_moves\ 5/a\ 6 //    
+  |#m #HI * [#visited #vinv #finv >\ 5a href="cic:/matita/tutorial/chapter10/bisim_end.def(10)"\ 6bisim_end\ 5/a\ 6 //]
    #p #front_tl #visited #Hn * #u_visited #r_visited * #u_frontier #r_frontier 
    #disjoint
-   cut (∃w.(moves ? w e1 = \fst p) ∧ (moves ? w e2 = \snd p)) 
-    [@(r_frontier … (memb_hd … ))] #rp
-   cut (beqb (\snd (\fst p)) (\snd (\snd p)) = true)
+   cut (\ 5a title="exists" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6w.(\ 5a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"\ 6moves\ 5/a\ 6 ? w e1 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 p) \ 5a title="logical and" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter9/moves.fix(0,1,7)"\ 6moves\ 5/a\ 6 ? w e2 \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 p)) 
+    [@(r_frontier … (\ 5a href="cic:/matita/tutorial/chapter5/memb_hd.def(5)"\ 6memb_hd\ 5/a\ 6 … ))] #rp
+   cut (\ 5a href="cic:/matita/tutorial/chapter4/beqb.def(2)"\ 6beqb\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 p)) (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 p)) \ 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)
     [cases rp #w * #fstp #sndp <fstp <sndp @(\b ?) 
-     @(proj1 … (equiv_sem … )) @same] #ptest 
-   >(bisim_step_true … ptest) @HI -HI 
-     [<plus_n_Sm //
-     |% [whd in ⊢ (??%?); >(disjoint … (memb_hd …)) whd in ⊢ (??%?); //
-        |#p1 #H (cases (orb_true_l … H)) [#eqp >(\P eqp) // |@r_visited]
+     @(\ 5a href="cic:/matita/basics/logic/proj1.def(2)"\ 6proj1\ 5/a\ 6 … (\ 5a href="cic:/matita/tutorial/chapter10/equiv_sem.def(16)"\ 6equiv_sem\ 5/a\ 6 … )) @same] #ptest 
+   >(\ 5a href="cic:/matita/tutorial/chapter10/bisim_step_true.def(10)"\ 6bisim_step_true\ 5/a\ 6 … ptest) @HI -HI 
+     [<\ 5a href="cic:/matita/arithmetics/nat/plus_n_Sm.def(4)"\ 6plus_n_Sm\ 5/a\ 6 //
+     |% [whd in ⊢ (??%?); >(disjoint … (\ 5a href="cic:/matita/tutorial/chapter5/memb_hd.def(5)"\ 6memb_hd\ 5/a\ 6 …)) whd in ⊢ (??%?); //
+        |#p1 #H (cases (\ 5a href="cic:/matita/basics/bool/orb_true_l.def(2)"\ 6orb_true_l\ 5/a\ 6 … H)) [#eqp >(\P eqp) // |@r_visited]
         ]
-     |whd % [@unique_append_unique @(andb_true_r … u_frontier)]
-      @unique_append_elim #q #H
-       [cases (memb_sons … (memb_filter_memb … H)) -H
-        #a * #m1 #m2 cases rp #w1 * #mw1 #mw2 @(ex_intro … (w1@[a]))
-        >moves_left >moves_left >mw1 >mw2 >m1 >m2 % // 
-       |@r_frontier @memb_cons //
+     |whd % [@\ 5a href="cic:/matita/tutorial/chapter5/unique_append_unique.def(6)"\ 6unique_append_unique\ 5/a\ 6 @(\ 5a href="cic:/matita/basics/bool/andb_true_r.def(4)"\ 6andb_true_r\ 5/a\ 6 … u_frontier)]
+      @\ 5a href="cic:/matita/tutorial/chapter5/unique_append_elim.dec"\ 6unique_append_elim\ 5/a\ 6 #q #H
+       [cases (\ 5a href="cic:/matita/tutorial/chapter10/memb_sons.def(8)"\ 6memb_sons\ 5/a\ 6 … (\ 5a href="cic:/matita/tutorial/chapter5/memb_filter_memb.def(5)"\ 6memb_filter_memb\ 5/a\ 6 … H)) -H
+        #a * #m1 #m2 cases rp #w1 * #mw1 #mw2 @(\ 5a href="cic:/matita/basics/logic/ex.con(0,1,2)"\ 6ex_intro\ 5/a\ 6 … (w1\ 5a title="append" href="cic:/fakeuri.def(1)"\ 6@\ 5/a\ 6\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6a]))
+        >\ 5a href="cic:/matita/tutorial/chapter9/moves_left.def(9)"\ 6moves_left\ 5/a\ 6 >\ 5a href="cic:/matita/tutorial/chapter9/moves_left.def(9)"\ 6moves_left\ 5/a\ 6 >mw1 >mw2 >m1 >m2 % // 
+       |@r_frontier @\ 5a href="cic:/matita/tutorial/chapter5/memb_cons.def(5)"\ 6memb_cons\ 5/a\ 6 //
        ]
-     |@unique_append_elim #q #H
-       [@injective_notb @(filter_true … H)
-       |cut ((q==p) = false
-         [|#Hpq whd in ⊢ (??%?); >Hpq @disjoint @memb_cons //]
-        cases (andb_true … u_frontier) #notp #_ @(\bf ?) 
-        @(not_to_not … not_eq_true_false) #eqqp <notp <eqqp >H //
+     |@\ 5a href="cic:/matita/tutorial/chapter5/unique_append_elim.dec"\ 6unique_append_elim\ 5/a\ 6 #q #H
+       [@\ 5a href="cic:/matita/basics/bool/injective_notb.def(4)"\ 6injective_notb\ 5/a\ 6 @(\ 5a href="cic:/matita/tutorial/chapter5/filter_true.def(5)"\ 6filter_true\ 5/a\ 6 … H)
+       |cut ((q\ 5a title="eqb" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6=p) \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 \ 5a href="cic:/matita/basics/bool/bool.con(0,2,0)"\ 6false\ 5/a\ 6
+         [|#Hpq whd in ⊢ (??%?); >Hpq @disjoint @\ 5a href="cic:/matita/tutorial/chapter5/memb_cons.def(5)"\ 6memb_cons\ 5/a\ 6 //]
+        cases (\ 5a href="cic:/matita/basics/bool/andb_true.def(5)"\ 6andb_true\ 5/a\ 6 … u_frontier) #notp #_ @(\bf ?) 
+        @(\ 5a href="cic:/matita/basics/logic/not_to_not.def(3)"\ 6not_to_not\ 5/a\ 6 … \ 5a href="cic:/matita/basics/bool/not_eq_true_false.def(3)"\ 6not_eq_true_false\ 5/a\ 6) #eqqp <notp <eqqp >H //
        ]
      ]
    ]  
@@ -279,60 +279,60 @@ 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:(pre S) × (pre S). memb ? p l = true → 
-  (beqb (\snd (\fst p)) (\snd (\snd p)) = true).
+definition all_true ≝ λS.λl.∀p:(\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S) \ 5a title="Product" href="cic:/fakeuri.def(1)"\ 6×\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S). \ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 ? p l \ 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 href="cic:/matita/tutorial/chapter4/beqb.def(2)"\ 6beqb\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 p)) (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 p)) \ 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).
 
-definition sub_sons ≝ λS,l,l1,l2.∀x:(pre S) × (pre S). 
-memb ? x l1 = true → sublist ? (sons ? l x) l2. 
+definition sub_sons ≝ λS,l,l1,l2.∀x:(\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S) \ 5a title="Product" href="cic:/fakeuri.def(1)"\ 6×\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter7/pre.def(1)"\ 6pre\ 5/a\ 6 S). 
+\ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 ? x l1 \ 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 href="cic:/matita/tutorial/chapter5/sublist.def(5)"\ 6sublist\ 5/a\ 6 ? (\ 5a href="cic:/matita/tutorial/chapter10/sons.def(7)"\ 6sons\ 5/a\ 6 ? l x) l2. 
 
 lemma bisim_complete: 
- ∀S,l,n.∀frontier,visited,visited_res:list ?.
all_true S visited →
sub_sons S l visited (frontier@visited) →
bisim S l n frontier visited = 〈true,visited_res〉 →
is_bisim S visited_res l ∧ sublist ? (frontier@visited) visited_res. 
+ ∀S,l,n.∀frontier,visited,visited_res:\ 5a href="cic:/matita/basics/list/list.ind(1,0,1)"\ 6list\ 5/a\ 6 ?.
\ 5a href="cic:/matita/tutorial/chapter10/all_true.def(8)"\ 6all_true\ 5/a\ 6 S visited →
\ 5a href="cic:/matita/tutorial/chapter10/sub_sons.def(8)"\ 6sub_sons\ 5/a\ 6 S l visited (frontier\ 5a title="append" href="cic:/fakeuri.def(1)"\ 6@\ 5/a\ 6visited) →
\ 5a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"\ 6bisim\ 5/a\ 6 S l n frontier visited \ 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 href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6,visited_res〉 →
\ 5a href="cic:/matita/tutorial/chapter10/is_bisim.def(8)"\ 6is_bisim\ 5/a\ 6 S visited_res l \ 5a title="logical and" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter5/sublist.def(5)"\ 6sublist\ 5/a\ 6 ? (frontier\ 5a title="append" href="cic:/fakeuri.def(1)"\ 6@\ 5/a\ 6visited) visited_res. 
 #S #l #n elim n
-  [#fron #vis #vis_res #_ #_ >bisim_never #H destruct
+  [#fron #vis #vis_res #_ #_ >\ 5a href="cic:/matita/tutorial/chapter10/bisim_never.def(10)"\ 6bisim_never\ 5/a\ 6 #H destruct
   |#m #Hind * 
     [(* case empty frontier *)
      -Hind #vis #vis_res #allv #H normalize in  ⊢ (%→?);
      #H1 destruct % #p 
       [#membp % [@(\P ?) @allv //| @H //]|#H1 @H1]
-    |#hd cases (true_or_false (beqb (\snd (\fst hd)) (\snd (\snd hd))))
+    |#hd cases (\ 5a href="cic:/matita/basics/bool/true_or_false.def(1)"\ 6true_or_false\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter4/beqb.def(2)"\ 6beqb\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 hd)) (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 hd))))
       [|(* case head of the frontier is non ok (absurd) *)
-       #H #tl #vis #vis_res #allv >(bisim_step_false … H) #_ #H1 destruct]
+       #H #tl #vis #vis_res #allv >(\ 5a href="cic:/matita/tutorial/chapter10/bisim_step_false.def(10)"\ 6bisim_step_false\ 5/a\ 6 … H) #_ #H1 destruct]
      (* frontier = hd:: tl and hd is ok *)
-     #H #tl #visited #visited_res #allv >(bisim_step_true … H)
+     #H #tl #visited #visited_res #allv >(\ 5a href="cic:/matita/tutorial/chapter10/bisim_step_true.def(10)"\ 6bisim_step_true\ 5/a\ 6 … H)
      (* new_visited = hd::visited are all ok *)
-     cut (all_true S (hd::visited)) 
-      [#p #H1 cases (orb_true_l … H1) [#eqp >(\P eqp) @H |@allv]]
+     cut (\ 5a href="cic:/matita/tutorial/chapter10/all_true.def(8)"\ 6all_true\ 5/a\ 6 S (hd\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:visited)) 
+      [#p #H1 cases (\ 5a href="cic:/matita/basics/bool/orb_true_l.def(2)"\ 6orb_true_l\ 5/a\ 6 … H1) [#eqp >(\P eqp) @H |@allv]]
      (* we now exploit the induction hypothesis *)
      #allh #subH #bisim cases (Hind … allh … bisim) -bisim -Hind
-      [#H1 #H2 % // #p #membp @H2 -H2 cases (memb_append … membp) -membp #membp
-        [cases (orb_true_l … membp) -membp #membp
-          [@memb_append_l2 >(\P membp) @memb_hd
-          |@memb_append_l1 @sublist_unique_append_l2 // 
+      [#H1 #H2 % // #p #membp @H2 -H2 cases (\ 5a href="cic:/matita/tutorial/chapter5/memb_append.def(5)"\ 6memb_append\ 5/a\ 6 … membp) -membp #membp
+        [cases (\ 5a href="cic:/matita/basics/bool/orb_true_l.def(2)"\ 6orb_true_l\ 5/a\ 6 … membp) -membp #membp
+          [@\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l2.def(5)"\ 6memb_append_l2\ 5/a\ 6 >(\P membp) @\ 5a href="cic:/matita/tutorial/chapter5/memb_hd.def(5)"\ 6memb_hd\ 5/a\ 6
+          |@\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l1.def(5)"\ 6memb_append_l1\ 5/a\ 6 @\ 5a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l2.def(6)"\ 6sublist_unique_append_l2\ 5/a\ 6 // 
           ]
-        |@memb_append_l2 @memb_cons //
+        |@\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l2.def(5)"\ 6memb_append_l2\ 5/a\ 6 @\ 5a href="cic:/matita/tutorial/chapter5/memb_cons.def(5)"\ 6memb_cons\ 5/a\ 6 //
         ] 
       |(* the only thing left to prove is the sub_sons invariant *)  
-     #x #membx cases (orb_true_l … membx)
+     #x #membx cases (\ 5a href="cic:/matita/basics/bool/orb_true_l.def(2)"\ 6orb_true_l\ 5/a\ 6 … membx)
       [(* case x = hd *) 
        #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 (true_or_false … (memb ? xa (x::visited)))
-        [(* xa visited - trivial *) #membxa @memb_append_l2 //
-        |(* xa new *) #membxa @memb_append_l1 @sublist_unique_append_l1 @memb_filter_l
+       cases (\ 5a href="cic:/matita/basics/bool/true_or_false.def(1)"\ 6true_or_false\ 5/a\ 6 … (\ 5a href="cic:/matita/tutorial/chapter5/memb.fix(0,2,4)"\ 6memb\ 5/a\ 6 ? xa (x\ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6:\ 5/a\ 6:visited)))
+        [(* xa visited - trivial *) #membxa @\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l2.def(5)"\ 6memb_append_l2\ 5/a\ 6 //
+        |(* xa new *) #membxa @\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l1.def(5)"\ 6memb_append_l1\ 5/a\ 6 @\ 5a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l1.def(6)"\ 6sublist_unique_append_l1\ 5/a\ 6 @\ 5a href="cic:/matita/tutorial/chapter5/memb_filter_l.def(5)"\ 6memb_filter_l\ 5/a\ 6
           [>membxa //|//]
         ]
       |(* case x in visited *)
-       #H1 #xa #membxa cases (memb_append … (subH x … H1 … membxa))  
-        [#H2 (cases (orb_true_l … H2)) 
-          [#H3 @memb_append_l2 <(\P H3) @memb_hd
-          |#H3 @memb_append_l1 @sublist_unique_append_l2 @H3
+       #H1 #xa #membxa cases (\ 5a href="cic:/matita/tutorial/chapter5/memb_append.def(5)"\ 6memb_append\ 5/a\ 6 … (subH x … H1 … membxa))  
+        [#H2 (cases (\ 5a href="cic:/matita/basics/bool/orb_true_l.def(2)"\ 6orb_true_l\ 5/a\ 6 … H2)) 
+          [#H3 @\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l2.def(5)"\ 6memb_append_l2\ 5/a\ 6 <(\P H3) @\ 5a href="cic:/matita/tutorial/chapter5/memb_hd.def(5)"\ 6memb_hd\ 5/a\ 6
+          |#H3 @\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l1.def(5)"\ 6memb_append_l1\ 5/a\ 6 @\ 5a href="cic:/matita/tutorial/chapter5/sublist_unique_append_l2.def(6)"\ 6sublist_unique_append_l2\ 5/a\ 6 @H3
           ]
-        |#H2 @memb_append_l2 @memb_cons @H2
+        |#H2 @\ 5a href="cic:/matita/tutorial/chapter5/memb_append_l2.def(5)"\ 6memb_append_l2\ 5/a\ 6 @\ 5a href="cic:/matita/tutorial/chapter5/memb_cons.def(5)"\ 6memb_cons\ 5/a\ 6 @H2
         ]
       ]
     ]
@@ -343,57 +343,56 @@ qed.
 prove that two expressions are equivalente if and only if they define
 the same language. *)
 
-definition equiv ≝ λSig.λre1,re2:re Sig. 
-  let e1 ≝ •(blank ? re1) in
-  let e2 ≝ •(blank ? re2) in
-  let n ≝ S (length ? (space_enum Sig (|\fst e1|) (|\fst e2|))) in
-  let sig ≝ (occ Sig e1 e2) in
-  (bisim ? sig n [〈e1,e2〉] []).
+definition equiv ≝ λSig.λre1,re2:\ 5a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"\ 6re\ 5/a\ 6 Sig. 
+  let e1 ≝ \ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"\ 6blank\ 5/a\ 6 ? re1) in
+  let e2 ≝ \ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"\ 6blank\ 5/a\ 6 ? re2) in
+  let n ≝ \ 5a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"\ 6S\ 5/a\ 6 (\ 5a href="cic:/matita/basics/list/length.fix(0,1,1)"\ 6length\ 5/a\ 6 ? (\ 5a href="cic:/matita/tutorial/chapter10/space_enum.def(5)"\ 6space_enum\ 5/a\ 6 Sig (\ 5a title="forget" 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="forget" href="cic:/fakeuri.def(1)"\ 6|\ 5/a\ 6\ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 e2|))) in
+  let sig ≝ (\ 5a href="cic:/matita/tutorial/chapter10/occ.def(7)"\ 6occ\ 5/a\ 6 Sig e1 e2) in
+  (\ 5a href="cic:/matita/tutorial/chapter10/bisim.fix(0,2,8)"\ 6bisim\ 5/a\ 6 ? sig n \ 5a title="cons" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6\ 5a title="Pair construction" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6e1,e2〉] \ 5a title="nil" href="cic:/fakeuri.def(1)"\ 6[\ 5/a\ 6]).
 
-theorem euqiv_sem : ∀Sig.∀e1,e2:re Sig.
-   \fst (equiv ? e1 e2) = true ↔ \sem{e1} =1 \sem{e2}.
+theorem euqiv_sem : ∀Sig.∀e1,e2:\ 5a href="cic:/matita/tutorial/chapter7/re.ind(1,0,1)"\ 6re\ 5/a\ 6 Sig.
+   \ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter10/equiv.def(9)"\ 6equiv\ 5/a\ 6 ? e1 e2) \ 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="iff" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 \ 5a title="in_l" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{e1} \ 5a title="extensional equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6\ 5a title="in_l" href="cic:/fakeuri.def(1)"\ 6\sem\ 5/a\ 6{e2}.
 #Sig #re1 #re2 %
-  [#H @eqP_trans [|@eqP_sym @re_embedding] @eqP_trans [||@re_embedding]
-   cut (equiv ? re1 re2 = 〈true,\snd (equiv ? re1 re2)〉)
+  [#H @\ 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/chapter8/re_embedding.def(13)"\ 6re_embedding\ 5/a\ 6] @\ 5a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"\ 6eqP_trans\ 5/a\ 6 [||@\ 5a href="cic:/matita/tutorial/chapter8/re_embedding.def(13)"\ 6re_embedding\ 5/a\ 6]
+   cut (\ 5a href="cic:/matita/tutorial/chapter10/equiv.def(9)"\ 6equiv\ 5/a\ 6 ? re1 re2 \ 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 href="cic:/matita/basics/bool/bool.con(0,1,0)"\ 6true\ 5/a\ 6,\ 5a title="pair pi2" href="cic:/fakeuri.def(1)"\ 6\snd\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter10/equiv.def(9)"\ 6equiv\ 5/a\ 6 ? re1 re2)〉)
      [<H //] #Hcut
-   cases (bisim_complete … Hcut) 
-     [2,3: #p whd in ⊢ ((??%?)→?); #abs @False_ind /2/] 
-   #Hbisim #Hsub @(bisim_to_sem … Hbisim) 
-   @Hsub @memb_hd
-  |#H @(bisim_correct ? (•(blank ? re1)) (•(blank ? re2))) 
-    [@eqP_trans [|@re_embedding] @eqP_trans [|@H] @eqP_sym @re_embedding
+   cases (\ 5a href="cic:/matita/tutorial/chapter10/bisim_complete.def(11)"\ 6bisim_complete\ 5/a\ 6 … Hcut) 
+     [2,3: #p whd in ⊢ ((??%?)→?); #abs @\ 5a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"\ 6False_ind\ 5/a\ 6 /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/absurd.def(2)"\ 6absurd\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/] 
+   #Hbisim #Hsub @(\ 5a href="cic:/matita/tutorial/chapter10/bisim_to_sem.def(18)"\ 6bisim_to_sem\ 5/a\ 6 … Hbisim) 
+   @Hsub @\ 5a href="cic:/matita/tutorial/chapter5/memb_hd.def(5)"\ 6memb_hd\ 5/a\ 6
+  |#H @(\ 5a href="cic:/matita/tutorial/chapter10/bisim_correct.def(17)"\ 6bisim_correct\ 5/a\ 6 ? (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"\ 6blank\ 5/a\ 6 ? re1)) (\ 5a title="eclose" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6(\ 5a href="cic:/matita/tutorial/chapter8/blank.fix(0,1,3)"\ 6blank\ 5/a\ 6 ? re2))) 
+    [@\ 5a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"\ 6eqP_trans\ 5/a\ 6 [|@\ 5a href="cic:/matita/tutorial/chapter8/re_embedding.def(13)"\ 6re_embedding\ 5/a\ 6] @\ 5a href="cic:/matita/tutorial/chapter4/eqP_trans.def(3)"\ 6eqP_trans\ 5/a\ 6 [|@H] @\ 5a href="cic:/matita/tutorial/chapter4/eqP_sym.def(3)"\ 6eqP_sym\ 5/a\ 6 @\ 5a href="cic:/matita/tutorial/chapter8/re_embedding.def(13)"\ 6re_embedding\ 5/a\ 6
     |// 
-    |% // #p whd in ⊢ ((??%?)→?); #abs @False_ind /2/  
-    |% // #p #H >(memb_single … H) @(ex_intro … ϵ) /2/
+    |% // #p whd in ⊢ ((??%?)→?); #abs @\ 5a href="cic:/matita/basics/logic/False_ind.fix(0,1,1)"\ 6False_ind\ 5/a\ 6 /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/absurd.def(2)"\ 6absurd\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/  
+    |% // #p #H >(\ 5a href="cic:/matita/tutorial/chapter5/memb_single.def(5)"\ 6memb_single\ 5/a\ 6 … H) @(\ 5a href="cic:/matita/basics/logic/ex.con(0,1,2)"\ 6ex_intro\ 5/a\ 6 … \ 5a title="epsilon" href="cic:/fakeuri.def(1)"\ 6ϵ\ 5/a\ 6) /\ 5span class="autotactic"\ 62\ 5span class="autotrace"\ 6 trace \ 5a href="cic:/matita/basics/logic/And.con(0,1,2)"\ 6conj\ 5/a\ 6\ 5/span\ 6\ 5/span\ 6/
     |#p #_ normalize //
     ]
   ]
 qed.
 
-lemma eqbnat_true : ∀n,m. eqbnat n m = true ↔ n = m.
-#n #m % [@eqbnat_true_to_eq | @eq_to_eqbnat_true]
+definition eqbnat ≝ λn,m:\ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6\ 5a href="cic:/matita/arithmetics/nat/eqb.fix(0,0,1)"\ 6eqb\ 5/a\ 6 n m.
+
+lemma eqbnat_true : ∀n,m. \ 5a href="cic:/matita/tutorial/chapter10/eqbnat.def(2)"\ 6eqbnat\ 5/a\ 6 n m \ 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="iff" href="cic:/fakeuri.def(1)"\ 6\ 5/a\ 6 n \ 5a title="leibnitz's equality" href="cic:/fakeuri.def(1)"\ 6=\ 5/a\ 6 m.
+#n #m % [@\ 5a href="cic:/matita/arithmetics/nat/eqb_true_to_eq.def(6)"\ 6eqb_true_to_eq\ 5/a\ 6 | @\ 5a href="cic:/matita/arithmetics/nat/eq_to_eqb_true.def(5)"\ 6eq_to_eqb_true\ 5/a\ 6]
 qed.
 
-definition DeqNat ≝ mk_DeqSet nat eqbnat eqbnat_true.
+definition DeqNat ≝ \ 5a href="cic:/matita/tutorial/chapter4/DeqSet.con(0,1,0)"\ 6mk_DeqSet\ 5/a\ 6 \ 5a href="cic:/matita/arithmetics/nat/nat.ind(1,0,0)"\ 6nat\ 5/a\ 6 \ 5a href="cic:/matita/arithmetics/nat/eqb.fix(0,0,1)"\ 6eqb\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter10/eqbnat_true.def(7)"\ 6eqbnat_true\ 5/a\ 6.
 
-definition a ≝ s DeqNat O.
-definition b ≝ s DeqNat (S O).
-definition c ≝ s DeqNat (S (S O)).
+definition a ≝ \ 5a href="cic:/matita/tutorial/chapter7/re.con(0,3,1)"\ 6s\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter10/DeqNat.def(8)"\ 6DeqNat\ 5/a\ 6 \ 5a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"\ 6O\ 5/a\ 6.
+definition b ≝ \ 5a href="cic:/matita/tutorial/chapter7/re.con(0,3,1)"\ 6s\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter10/DeqNat.def(8)"\ 6DeqNat\ 5/a\ 6 (\ 5a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"\ 6S\ 5/a\ 6 \ 5a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"\ 6O\ 5/a\ 6).
+definition c ≝ \ 5a href="cic:/matita/tutorial/chapter7/re.con(0,3,1)"\ 6s\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter10/DeqNat.def(8)"\ 6DeqNat\ 5/a\ 6 (\ 5a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"\ 6S\ 5/a\ 6 (\ 5a href="cic:/matita/arithmetics/nat/nat.con(0,2,0)"\ 6S\ 5/a\ 6 \ 5a href="cic:/matita/arithmetics/nat/nat.con(0,1,0)"\ 6O\ 5/a\ 6)).
 
-definition exp1 ≝ ((a·b)^*·a).
-definition exp2 ≝ a·(b·a)^*.
-definition exp4 ≝ (b·a)^*.
+definition exp1 ≝ ((\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/b.def(9)"\ 6b\ 5/a\ 6)\ 5a title="re star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6).
+definition exp2 ≝ \ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6(\ 5a href="cic:/matita/tutorial/chapter10/b.def(9)"\ 6b\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6)\ 5a title="re star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*.
+definition exp4 ≝ (\ 5a href="cic:/matita/tutorial/chapter10/b.def(9)"\ 6b\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6)\ 5a title="re star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*.
 
-definition exp6 ≝ a·(a ·a ·b^* + b^* ).
-definition exp7 ≝ a · a^* · b^*.
+definition exp6 ≝ \ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6(\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6 \ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6 \ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/b.def(9)"\ 6b\ 5/a\ 6\ 5a title="re star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6\ 5a title="re or" href="cic:/fakeuri.def(1)"\ 6+\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter10/b.def(9)"\ 6b\ 5/a\ 6\ 5a title="re star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6* ).
+definition exp7 ≝ \ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6 \ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter10/b.def(9)"\ 6b\ 5/a\ 6\ 5a title="re star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*.
 
-definition exp8 ≝ a·a·a·a·a·a·a·a·(a^* ).
-definition exp9 ≝ (a·a·a + a·a·a·a·a)^*.
+definition exp8 ≝ \ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6(\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6* ).
+definition exp9 ≝ (\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6 \ 5a title="re or" href="cic:/fakeuri.def(1)"\ 6+\ 5/a\ 6 \ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6\ 5a title="re cat" href="cic:/fakeuri.def(1)"\ 6·\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/a.def(9)"\ 6a\ 5/a\ 6)\ 5a title="re star" href="cic:/fakeuri.def(1)"\ 6^\ 5/a\ 6*.
 
-example ex1 : \fst (equiv ? (exp8+exp9) exp9) = true.
+example ex1 : \ 5a title="pair pi1" href="cic:/fakeuri.def(1)"\ 6\fst\ 5/a\ 6 (\ 5a href="cic:/matita/tutorial/chapter10/equiv.def(9)"\ 6equiv\ 5/a\ 6 ? (\ 5a href="cic:/matita/tutorial/chapter10/exp8.def(10)"\ 6exp8\ 5/a\ 6\ 5a title="re or" href="cic:/fakeuri.def(1)"\ 6+\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/exp9.def(10)"\ 6exp9\ 5/a\ 6\ 5a href="cic:/matita/tutorial/chapter10/exp9.def(10)"\ 6exp9\ 5/a\ 6\ 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.
 normalize // qed.
 
 
-
-
-