X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fgrammar%2Ftstc.ma;h=a861d18a6f55f5a25d42be2cc001ea21fd3bb7fa;hb=52e675f555f559c047d5449db7fc89a51b977d35;hp=defc91ed95e244b5040bc87b939e6d1104373986;hpb=75fac6d60f67a4dfa38ea6c2cc45a18eda5d8996;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/tstc.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/tstc.ma index defc91ed9..a861d18a6 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/tstc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/tstc.ma @@ -12,97 +12,97 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/iso_2.ma". +include "basic_2/notation/relations/topiso_2.ma". include "basic_2/grammar/term_simple.ma". (* SAME TOP TERM CONSTRUCTOR ************************************************) inductive tstc: relation term ≝ | tstc_atom: ∀I. tstc (⓪{I}) (⓪{I}) - | tstc_pair: ∀I,V1,V2,T1,T2. tstc (②{I} V1. T1) (②{I} V2. T2) + | tstc_pair: ∀I,V1,V2,T1,T2. tstc (②{I}V1.T1) (②{I}V2.T2) . -interpretation "same top constructor (term)" 'Iso T1 T2 = (tstc T1 T2). +interpretation "same top constructor (term)" 'TopIso T1 T2 = (tstc T1 T2). (* Basic inversion lemmas ***************************************************) -fact tstc_inv_atom1_aux: ∀T1,T2. T1 ≃ T2 → ∀I. T1 = ⓪{I} → T2 = ⓪{I}. +fact tstc_inv_atom1_aux: ∀T1,T2. T1 ≂ T2 → ∀I. T1 = ⓪{I} → T2 = ⓪{I}. #T1 #T2 * -T1 -T2 // #J #V1 #V2 #T1 #T2 #I #H destruct -qed. +qed-. (* Basic_1: was: iso_gen_sort iso_gen_lref *) -lemma tstc_inv_atom1: ∀I,T2. ⓪{I} ≃ T2 → T2 = ⓪{I}. -/2 width=3/ qed-. +lemma tstc_inv_atom1: ∀I,T2. ⓪{I} ≂ T2 → T2 = ⓪{I}. +/2 width=3 by tstc_inv_atom1_aux/ qed-. -fact tstc_inv_pair1_aux: ∀T1,T2. T1 ≃ T2 → ∀I,W1,U1. T1 = ②{I}W1.U1 → +fact tstc_inv_pair1_aux: ∀T1,T2. T1 ≂ T2 → ∀I,W1,U1. T1 = ②{I}W1.U1 → ∃∃W2,U2. T2 = ②{I}W2. U2. #T1 #T2 * -T1 -T2 [ #J #I #W1 #U1 #H destruct -| #J #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct /2 width=3/ +| #J #V1 #V2 #T1 #T2 #I #W1 #U1 #H destruct /2 width=3 by ex1_2_intro/ ] -qed. +qed-. (* Basic_1: was: iso_gen_head *) -lemma tstc_inv_pair1: ∀I,W1,U1,T2. ②{I}W1.U1 ≃ T2 → +lemma tstc_inv_pair1: ∀I,W1,U1,T2. ②{I}W1.U1 ≂ T2 → ∃∃W2,U2. T2 = ②{I}W2. U2. -/2 width=5/ qed-. +/2 width=5 by tstc_inv_pair1_aux/ qed-. -fact tstc_inv_atom2_aux: ∀T1,T2. T1 ≃ T2 → ∀I. T2 = ⓪{I} → T1 = ⓪{I}. +fact tstc_inv_atom2_aux: ∀T1,T2. T1 ≂ T2 → ∀I. T2 = ⓪{I} → T1 = ⓪{I}. #T1 #T2 * -T1 -T2 // #J #V1 #V2 #T1 #T2 #I #H destruct -qed. +qed-. -lemma tstc_inv_atom2: ∀I,T1. T1 ≃ ⓪{I} → T1 = ⓪{I}. -/2 width=3/ qed-. +lemma tstc_inv_atom2: ∀I,T1. T1 ≂ ⓪{I} → T1 = ⓪{I}. +/2 width=3 by tstc_inv_atom2_aux/ qed-. -fact tstc_inv_pair2_aux: ∀T1,T2. T1 ≃ T2 → ∀I,W2,U2. T2 = ②{I}W2.U2 → - ∃∃W1,U1. T1 = ②{I}W1. U1. +fact tstc_inv_pair2_aux: ∀T1,T2. T1 ≂ T2 → ∀I,W2,U2. T2 = ②{I}W2.U2 → + ∃∃W1,U1. T1 = ②{I}W1.U1. #T1 #T2 * -T1 -T2 [ #J #I #W2 #U2 #H destruct -| #J #V1 #V2 #T1 #T2 #I #W2 #U2 #H destruct /2 width=3/ +| #J #V1 #V2 #T1 #T2 #I #W2 #U2 #H destruct /2 width=3 by ex1_2_intro/ ] -qed. +qed-. -lemma tstc_inv_pair2: ∀I,T1,W2,U2. T1 ≃ ②{I}W2.U2 → - ∃∃W1,U1. T1 = ②{I}W1. U1. -/2 width=5/ qed-. +lemma tstc_inv_pair2: ∀I,T1,W2,U2. T1 ≂ ②{I}W2.U2 → + ∃∃W1,U1. T1 = ②{I}W1.U1. +/2 width=5 by tstc_inv_pair2_aux/ qed-. (* Basic properties *********************************************************) (* Basic_1: was: iso_refl *) -lemma tstc_refl: ∀T. T ≃ T. +lemma tstc_refl: reflexive … tstc. #T elim T -T // qed. -lemma tstc_sym: ∀T1,T2. T1 ≃ T2 → T2 ≃ T1. +lemma tstc_sym: symmetric … tstc. #T1 #T2 #H elim H -T1 -T2 // -qed. +qed-. -lemma tstc_dec: ∀T1,T2. Decidable (T1 ≃ T2). +lemma tstc_dec: ∀T1,T2. Decidable (T1 ≂ T2). * #I1 [2: #V1 #T1 ] * #I2 [2,4: #V2 #T2 ] [ elim (eq_item2_dec I1 I2) #HI12 - [ destruct /2 width=1/ + [ destruct /2 width=1 by tstc_pair, or_introl/ | @or_intror #H - elim (tstc_inv_pair1 … H) -H #V #T #H destruct /2 width=1/ + elim (tstc_inv_pair1 … H) -H #V #T #H destruct /2 width=1 by/ ] | @or_intror #H lapply (tstc_inv_atom1 … H) -H #H destruct | @or_intror #H lapply (tstc_inv_atom2 … H) -H #H destruct | elim (eq_item0_dec I1 I2) #HI12 - [ destruct /2 width=1/ + [ destruct /2 width=1 by or_introl/ | @or_intror #H - lapply (tstc_inv_atom2 … H) -H #H destruct /2 width=1/ + lapply (tstc_inv_atom2 … H) -H #H destruct /2 width=1 by/ ] ] qed. -lemma simple_tstc_repl_dx: ∀T1,T2. T1 ≃ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄. +lemma simple_tstc_repl_dx: ∀T1,T2. T1 ≂ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄. #T1 #T2 * -T1 -T2 // #I #V1 #V2 #T1 #T2 #H elim (simple_inv_pair … H) -H #J #H destruct // -qed. (**) (* remove from index *) +qed-. -lemma simple_tstc_repl_sn: ∀T1,T2. T1 ≃ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄. -/3 width=3/ qed-. +lemma simple_tstc_repl_sn: ∀T1,T2. T1 ≂ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄. +/3 width=3 by simple_tstc_repl_dx, tstc_sym/ qed-.