X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fdemo%2Fpropositional_sequent_calculus.ma;h=fb5278d6bdd78d44646652e6111317dc0fe3e370;hb=dcb1f00a9377a496b65d34f77a9c860b506df6c5;hp=4498676a3ef10f8a793464138f94e2638e130730;hpb=af130d273b6be7fbcc2fb2504f3b28ef8fa2344f;p=helm.git diff --git a/helm/software/matita/library/demo/propositional_sequent_calculus.ma b/helm/software/matita/library/demo/propositional_sequent_calculus.ma index 4498676a3..fb5278d6b 100644 --- a/helm/software/matita/library/demo/propositional_sequent_calculus.ma +++ b/helm/software/matita/library/demo/propositional_sequent_calculus.ma @@ -163,11 +163,10 @@ let rec and_of_list l ≝ | cons F l' ⇒ FAnd F (and_of_list l') ]. -alias id "Nil" = "cic:/matita/list/list.ind#xpointer(1/1/1)". let rec or_of_list l ≝ match l with - [ Nil ⇒ FFalse - | Cons F l' ⇒ FOr F (or_of_list l') + [ nil ⇒ FFalse + | cons F l' ⇒ FOr F (or_of_list l') ]. definition formula_of_sequent ≝ @@ -397,7 +396,7 @@ theorem mem_to_exists_l1_l2: [ simplify in H1; destruct H1 | simplify in H2; - apply (bool_elim ? (eq n t)); + apply (bool_elim ? (eq n a)); intro; [ apply (ex_intro ? ? []); apply (ex_intro ? ? l1); @@ -409,8 +408,8 @@ theorem mem_to_exists_l1_l2: elim (H H1 H2); elim H4; rewrite > H5; - apply (ex_intro ? ? (t::a)); - apply (ex_intro ? ? a1); + apply (ex_intro ? ? (a::a1)); + apply (ex_intro ? ? a2); simplify; reflexivity ] @@ -423,8 +422,7 @@ lemma same_atom_to_eq: ∀f1,f2. same_atom f1 f2 = true → f1=f2. [1,2: simplify in H; destruct H - | generalize in match H; clear H; - elim f2; + | elim f2 in H ⊢ %; [1,2: simplify in H; destruct H @@ -434,7 +432,7 @@ lemma same_atom_to_eq: ∀f1,f2. same_atom f1 f2 = true → f1=f2. |4,5: simplify in H2; destruct H2 - | simplify in H2; + | simplify in H1; destruct H1 ] |4,5: @@ -468,7 +466,7 @@ lemma mem_same_atom_to_exists: [ simplify in H; destruct H | simplify in H1; - apply (bool_elim ? (same_atom f t)); + apply (bool_elim ? (same_atom f a)); intros; [ elim (same_atom_to_exists ? ? H2); autobatch @@ -492,14 +490,14 @@ lemma look_for_axiom: simplify; reflexivity | intros; - generalize in match (refl_eq ? (mem ? same_atom t l2)); - elim (mem ? same_atom t l2) in ⊢ (? ? ? %→?); + generalize in match (refl_eq ? (mem ? same_atom a l2)); + elim (mem ? same_atom a l2) in ⊢ (? ? ? %→?); [ left; elim (mem_to_exists_l1_l2 ? ? ? ? same_atom_to_eq H1); elim H2; clear H2; elim (mem_same_atom_to_exists ? ? H1); rewrite > H2 in H3; - apply (ex_intro ? ? a2); + apply (ex_intro ? ? a3); rewrite > H2; apply (ex_intro ? ? []); simplify; @@ -507,22 +505,22 @@ lemma look_for_axiom: | elim (H l2); [ left; decompose; - apply (ex_intro ? ? a); - apply (ex_intro ? ? (t::a1)); + apply (ex_intro ? ? a1); + apply (ex_intro ? ? (a::a2)); simplify; - apply (ex_intro ? ? a2); apply (ex_intro ? ? a3); + apply (ex_intro ? ? a4); autobatch | right; intro; - apply (bool_elim ? (same_atom t (FAtom n1))); + apply (bool_elim ? (same_atom a (FAtom n1))); [ intro; rewrite > (eq_to_eq_mem ? ? transitiveb_same_atom ? ? ? H3) in H1; rewrite > H1; autobatch | intro; change in ⊢ (? ? (? % ?) ?) with - (match same_atom (FAtom n1) t with + (match same_atom (FAtom n1) a with [true ⇒ true |false ⇒ mem ? same_atom (FAtom n1) l ]); @@ -589,22 +587,16 @@ lemma sizel_0_no_axiom_is_tautology: intros; lapply (H1 (λn.mem ? same_atom (FAtom n) l1)); clear H1; simplify in Hletin; - generalize in match Hletin; clear Hletin; - generalize in match H2; clear H2; - generalize in match H; clear H; - elim l1 0; + elim l1 in Hletin H2 H ⊢ % 0; [ intros; simplify in H2; - generalize in match H2; clear H2; - generalize in match H1; clear H1; - generalize in match H; clear H; - elim l2 0; + elim l2 in H2 H1 H ⊢ % 0; [ intros; simplify in H2; destruct H2 | simplify; intro; - elim t; + elim a; [ right; apply (ex_intro ? ? []); simplify; @@ -617,10 +609,10 @@ lemma sizel_0_no_axiom_is_tautology: elim (not_eq_nil_append_cons ? ? ? ? H6) | elim H4; right; - apply (ex_intro ? ? (FFalse::a)); + apply (ex_intro ? ? (FFalse::a1)); simplify; elim H5; - apply (ex_intro ? ? a1); + apply (ex_intro ? ? a2); autobatch |3,4: autobatch | assumption @@ -632,7 +624,7 @@ lemma sizel_0_no_axiom_is_tautology: elim (not_eq_nil_append_cons ? ? ? ? H5) | right; elim H4; - apply (ex_intro ? ? (FAtom n::a)); + apply (ex_intro ? ? (FAtom n::a1)); simplify; elim H; autobatch @@ -645,11 +637,11 @@ lemma sizel_0_no_axiom_is_tautology: ] ] | intro; - elim t; + elim a; [ elim H; [ left; elim H4; - apply (ex_intro ? ? (FTrue::a)); + apply (ex_intro ? ? (FTrue::a1)); simplify; elim H5; autobatch @@ -669,7 +661,7 @@ lemma sizel_0_no_axiom_is_tautology: | elim H; [ left; elim H4; - apply (ex_intro ? ? (FAtom n::a)); + apply (ex_intro ? ? (FAtom n::a1)); simplify; elim H5; autobatch @@ -678,8 +670,7 @@ lemma sizel_0_no_axiom_is_tautology: | assumption | lapply (H2 n1); clear H2; simplify in Hletin; - generalize in match Hletin; clear Hletin; - elim (eqb n1 n); + elim (eqb n1 n) in Hletin ⊢ %; [ simplify in H2; rewrite > H2; autobatch @@ -694,10 +685,7 @@ lemma sizel_0_no_axiom_is_tautology: rewrite > eqb_n_n in H3; simplify in H3; (* - generalize in match H3; - generalize in match H1; clear H1; - generalize in match H; clear H; - elim l 0; + elim l in H3 H1 H ⊢ % 0; [ elim l2 0; [ intros; simplify in H2; @@ -715,9 +703,7 @@ lemma sizel_0_no_axiom_is_tautology: ] [ autobatch - | generalize in match H4; clear H4; - generalize in match H2; clear H2; - elim t1; + | elim t1 in H4 H2 ⊢ %; [ | | @@ -748,22 +734,22 @@ lemma completeness_base: elim S 1; clear S; simplify in ⊢ (?→%→?); intros; - elim (look_for_axiom t t1); + elim (look_for_axiom a b); [ decompose; rewrite > H2; clear H2; rewrite > H4; clear H4; - apply (exchangeL ? a1 a2 (FAtom a)); - apply (exchangeR ? a3 a4 (FAtom a)); - apply axiom - | elim (sizel_0_no_axiom_is_tautology t t1 H H1 H2); + apply (ExchangeL ? a2 a3 (FAtom a1)); + apply (ExchangeR ? a4 a5 (FAtom a1)); + apply Axiom + | elim (sizel_0_no_axiom_is_tautology a b H H1 H2); [ decompose; rewrite > H3; - apply (exchangeL ? a a1 FFalse); - apply falseL + apply (ExchangeL ? a1 a2 FFalse); + apply FalseL | decompose; rewrite > H3; - apply (exchangeR ? a a1 FTrue); - apply trueR + apply (ExchangeR ? a1 a2 FTrue); + apply TrueR ] ] qed.