X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Ftutorial%2Fchapter8.ma;h=427b148427f3e34efde54a55286a284424c0b1be;hb=5b93ea047903b606979705ed25a6df6504fd027c;hp=ca6b973c959b7b4ef40980d36f3bff9669a203e0;hpb=b58a13d78f5c7a37000538429aeefcd54662b570;p=helm.git diff --git a/matita/matita/lib/tutorial/chapter8.ma b/matita/matita/lib/tutorial/chapter8.ma index ca6b973c9..427b14842 100644 --- a/matita/matita/lib/tutorial/chapter8.ma +++ b/matita/matita/lib/tutorial/chapter8.ma @@ -270,8 +270,8 @@ lemma sem_star_w : ∀S.∀i:pitem S.∀w. (* Below are a few, simple, semantic properties of items. In particular: - not_epsilon_item : ∀S:DeqSet.∀i:pitem S. ¬ (\sem{i} ϵ). - epsilon_pre : ∀S.∀e:pre S. (\sem{i} ϵ) ↔ (\snd e = true). -- minus_eps_item: ∀S.∀i:pitem S. \sem{i} =1 \sem{i}-{[ ]}. -- minus_eps_pre: ∀S.∀e:pre S. \sem{\fst e} =1 \sem{e}-{[ ]}. +- minus_eps_item: ∀S.∀i:pitem S. \sem{i} ≐ \sem{i}-{[ ]}. +- minus_eps_pre: ∀S.∀e:pre S. \sem{\fst e} ≐ \sem{e}-{[ ]}. The first property is proved by a simple induction on $i$; the other results are easy corollaries. We need an auxiliary lemma first. *) @@ -297,14 +297,14 @@ lemma true_to_epsilon : ∀S.∀e:pre S. snd … e = true → ϵ ∈ e. #S * #i #b #btrue normalize in btrue; >btrue %2 // qed. -lemma minus_eps_item: ∀S.∀i:pitem S. \sem{i} ≐ \sem{i}-{[ ]}. +lemma minus_eps_item: ∀S.∀i:pitem S. \sem{i} ≐ \sem{i}-{ϵ}. #S #i #w % [#H whd % // normalize @(not_to_not … (not_epsilon_lp …i)) // |* // ] qed. -lemma minus_eps_pre: ∀S.∀e:pre S. \sem{fst ?? e} ≐ \sem{e}-{[ ]}. +lemma minus_eps_pre: ∀S.∀e:pre S. \sem{fst ?? e} ≐ \sem{e}-{ϵ}. #S * #i * [>sem_pre_true normalize in ⊢ (??%?); #w % [/3/ | * * // #H1 #H2 @False_ind @(absurd …H1 H2)]