projects
/
helm.git
/ commitdiff
commit
grep
author
committer
pickaxe
?
search:
re
summary
|
shortlog
|
log
|
commit
| commitdiff |
tree
raw
|
patch
| inline |
side by side
(parent:
5e6f74d
)
commit by user andrea
author
matitaweb
<claudio.sacerdoticoen@unibo.it>
Thu, 1 Mar 2012 13:00:50 +0000
(13:00 +0000)
committer
matitaweb
<claudio.sacerdoticoen@unibo.it>
Thu, 1 Mar 2012 13:00:50 +0000
(13:00 +0000)
weblib/tutorial/chapter6.ma
patch
|
blob
|
history
diff --git
a/weblib/tutorial/chapter6.ma
b/weblib/tutorial/chapter6.ma
index 430fc3f2140b002580629e498d0a89762b2a808d..cadc0065c3bc06a1b610a5312827fdf9a680ea8b 100644
(file)
--- a/
weblib/tutorial/chapter6.ma
+++ b/
weblib/tutorial/chapter6.ma
@@
-1,6
+1,4
@@
-include "basics/list.ma".
-include "basics/sets.ma".
-include "basics/deqsets.ma".
+include "tutorial/chapter5.ma".
definition word ≝ λS:DeqSet.list S.
@@
-129,4
+127,3
@@
lemma distr_cat_r_eps: ∀S.∀A,C:word S →Prop.
(A ∪ {ϵ}) · C =1 A · C ∪ C.
#S #A #C @eqP_trans [|@distr_cat_r |@eqP_union_l @epsilon_cat_l]
qed.
-