From 6a16a37b5b4cbea5f5216247182d5bb99a0d8d65 Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Thu, 6 Nov 2008 16:19:04 +0000 Subject: [PATCH] natural deduction support and example split, seems to almost work. --- .../matita/library/demo/natural_deduction.ma | 207 +++--------------- .../library/demo/natural_deduction_support.ma | 158 +++++++++++++ helm/software/matita/library/depends | 186 ++++++++-------- 3 files changed, 281 insertions(+), 270 deletions(-) create mode 100644 helm/software/matita/library/demo/natural_deduction_support.ma diff --git a/helm/software/matita/library/demo/natural_deduction.ma b/helm/software/matita/library/demo/natural_deduction.ma index 234b99af7..ba0f565e4 100644 --- a/helm/software/matita/library/demo/natural_deduction.ma +++ b/helm/software/matita/library/demo/natural_deduction.ma @@ -12,198 +12,49 @@ (* *) (**************************************************************************) -(*definition cast ≝ λA,B:CProp.λa:A.a.*) -axiom cast: ∀A,B:CProp.B → A. +include "demo/natural_deduction_support.ma". -(*notation < "\infrule (t\atop ⋮) (b \ALPOSTODI a) (? \ERROR)" with precedence 19 -for @{ 'caste $a $b $t }. -interpretation "cast" 'caste a b t = (cast a b t).*) -notation < "\infrule (t\atop ⋮) mstyle color #ff0000 (b) (? \ERROR)" with precedence 19 -for @{ 'caste $a $b $t }. -interpretation "cast" 'caste a b t = (cast a b t). +lemma ex1 : ΠA,B,C,E: CProp. -notation < "\infrule (t\atop ⋮) a ?" with precedence 19 for @{ 'cast $a $t }. -interpretation "cast" 'cast a t = (cast a a t). + (A ⇒ E) ∨ B ⇒ A ∧ C ⇒ (E ∧ C) ∨ B. -definition assumpt ≝ λA:CProp.λa:A.a. - -notation < "[ a ] \sup mstyle color #ff0000 (H)" with precedence 19 for @{ 'asse $a $H }. -interpretation "assumption" 'asse a H = (cast _ _ (assumpt a (cast _ _ H))). - -notation < "[ a ] \sup H" with precedence 19 for @{ 'ass $a $H }. -interpretation "assumption" 'ass a H = (cast a a (assumpt a (cast a a H))). - -inductive Imply (A,B:CProp) : CProp ≝ - Imply_intro: (A → B) → Imply A B. - -notation "hbox(a break ⇒ b)" right associative with precedence 20 for @{ 'Imply $a $b }. -interpretation "Imply" 'Imply a b = (Imply a b). - -notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000(ab) (⇒\sub\i \emsp ident H \ERROR) " with precedence 19 -for @{ 'Imply_introe $xxx $ab (λ${ident H}:$p.$b) }. -interpretation "Imply_intro" 'Imply_introe xxx ab \eta.b = (cast xxx ab (Imply_intro _ _ b)). - -notation < "\infrule hbox(\emsp b \emsp) ab (⇒\sub\i \emsp ident H) " with precedence 19 -for @{ 'Imply_intro $ab (λ${ident H}:$p.$b) }. -interpretation "Imply_intro" 'Imply_intro ab \eta.b = (cast ab ab (Imply_intro _ _ b)). - -definition Imply_elim ≝ λA,B.λf:Imply A B.λa:A.match f with [ Imply_intro g ⇒ g a]. - -notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b (⇒\sub\e) " with precedence 19 for @{ 'Imply_elim $ab $a $b }. -interpretation "Imply_elim" 'Imply_elim ab a b = (cast _ b (Imply_elim _ _ ab a)). - -inductive And (A,B:CProp) : CProp ≝ - And_intro: A → B → And A B. - -interpretation "constructive and" 'and x y = (And x y). - -notation < "\infrule hbox(\emsp a \emsp\emsp\emsp b \emsp) ab (∧\sub\i)" with precedence 19 for @{ 'And_intro $a $b $ab }. -interpretation "And_intro" 'And_intro a b ab = (cast _ ab (And_intro _ _ a b)). - -definition And_elim_l ≝ - λA,B.λc:A∧B.match c with [ And_intro a b ⇒ a ]. - -notation < "\infrule hbox(\emsp ab \emsp) a (∧\sub\e\sup\l)" with precedence 19 for @{ 'And_elim_l $ab $a }. -interpretation "And_elim_l" 'And_elim_l ab a = (cast _ a (And_elim_l _ _ ab)). - -definition And_elim_r ≝ - λA,B.λc:A∧B.match c with [ And_intro a b ⇒ b ]. - -notation < "\infrule hbox(\emsp ab \emsp) b (∧\sub\e\sup\r)" with precedence 19 for @{ 'And_elim_r $ab $b }. -interpretation "And_elim_r" 'And_elim_r ab b = (cast _ b (And_elim_r _ _ ab)). - -inductive Or (A,B:CProp) : CProp ≝ - | Or_intro_l: A → Or A B - | Or_intro_r: B → Or A B. + intros 4 (A B C E);apply (prove ((A⇒E)∨B⇒A∧C⇒E∧C∨B)); -interpretation "constructive or" 'or x y = (Or x y). - -notation < "\infrule hbox(\emsp a \emsp) ab (∨\sub\i\sup\l)" with precedence 19 for @{ 'Or_intro_l $a $ab }. -interpretation "Or_intro_l" 'Or_intro_l a ab = (cast _ ab (Or_intro_l _ _ a)). - -notation < "\infrule hbox(\emsp b \emsp) ab (∨\sub\i\sup\l)" with precedence 19 for @{ 'Or_intro_r $b $ab }. -interpretation "Or_intro_l" 'Or_intro_r b ab = (cast _ ab (Or_intro_r _ _ b)). - -definition Or_elim ≝ - λA,B,C:CProp.λc:A∨B.λfa: A → C.λfb: B → C. - match c with [ Or_intro_l a ⇒ fa a | Or_intro_r b ⇒ fb b]. - -notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp ac \emsp\emsp\emsp bc \emsp) c (∨\sub\e \emsp ident Ha \emsp ident Hb)" with precedence 19 -for @{ 'Or_elim $ab (λ${ident Ha}:$ta.$ac) (λ${ident Hb}:$tb.$bc) $c }. -interpretation "Or_elim" 'Or_elim ab \eta.ac \eta.bc c = (cast _ c (Or_elim _ _ _ ab ac bc)). - -inductive Exists (A:Type) (P:A→CProp) : CProp ≝ - Exists_intro: ∀w:A. P w → Exists A P. - -interpretation "constructive ex" 'exists \eta.x = (Exists _ x). - -notation < "\infrule hbox(\emsp Pn \emsp) Px (∃\sub\i)" with precedence 19 -for @{ 'Exists_intro $Pn $Px }. -interpretation "Exists_intro" 'Exists_intro Pn Px = (cast _ Px (Exists_intro _ _ _ Pn)). - -definition Exists_elim ≝ - λA:Type.λP:A→CProp.λC:CProp.λc:∃x:A.P x.λH:(∀x.P x → C). - match c with [ Exists_intro w p ⇒ H w p ]. - -notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) c (∃\sub\e \emsp ident n \emsp ident HPn)" with precedence 19 -for @{ 'Exists_elim $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }. -interpretation "Exists_elim" 'Exists_elim ExPx Pc c = (cast _ c (Exists_elim _ _ _ ExPx Pc)). - -inductive Forall (A:Type) (P:A→CProp) : CProp ≝ - Forall_intro: (∀n:A. P n) → Forall A P. - -notation "\forall ident x:A.break term 19 Px" with precedence 20 -for @{ 'Forall (λ${ident x}:$A.$Px) }. -interpretation "Forall" 'Forall \eta.Px = (Forall _ Px). - -notation < "\infrule hbox(\emsp Px \emsp) Pn (∀\sub\i \emsp ident x)" with precedence 19 -for @{ 'Forall_intro (λ${ident x}:$tx.$Px) $Pn }. -interpretation "Forall_intro" 'Forall_intro Px Pn = (cast _ Pn (Forall_intro _ _ Px)). - -definition Forall_elim ≝ - λA:Type.λP:A→CProp.λn:A.λf:∀x:A.P x.match f with [ Forall_intro g ⇒ g n ]. - -notation < "\infrule hbox(\emsp Px \emsp) Pn (∀\sub\i)" with precedence 19 for @{ 'Forall_elim $Px $Pn }. -interpretation "Forall_elim" 'Forall_elim Px Pn = (cast _ Pn (Forall_elim _ _ _ Px)). - -axiom A: CProp. -axiom B: CProp. -axiom C: CProp. -axiom D: CProp. -axiom E: CProp. - - -notation > "[H]" with precedence 90 -for @{ assumpt ? (cast ? ? $H)}. -notation > "⇒\sub\i [ident H] term 90 b" with precedence 19 -for @{ Imply_intro ?? (λ${ident H}.cast ? $b ?) }. -notation > "⇒\sub\e term 90 ab term 90 a" with precedence 19 -for @{ Imply_elim ?? (cast ? $ab ?) (cast $a $a ?) }. -notation > "∧\sub\i term 90 a term 90 b" with precedence 19 -for @{ And_intro ?? (cast ? $a ?) (cast ? $b ?) }. -(*notation > "∧\sub\e\sup\l term 90 ab" with precedence 19 -for @{ And_elim_l ?? (cast (? ∧ False) $ab ?) }. -notation > "∧\sub\e\sup\l term 90 a ∧ term 90 b" with precedence 19 -for @{ And_elim_l ?? (cast (? ∧ $b) ($a ∧ $b) ?) }. *) -notation > "∧\sub\e\sup\l term 90 ab" with precedence 19 -for @{ And_elim_l ?? (cast $ab $ab ?) }. (* CSC: WRONG *) -notation > "∧\sub\e\sup\r term 90 ab" with precedence 19 -for @{ And_elim_r ?? (cast $ab $ab ?) }. (* CSC: WRONG *) -notation > "∨\sub\i\sup\l term 90 a" with precedence 19 -for @{ Or_intro_l ?? (cast ? $a ?) }. -notation > "∨\sub\i\sup\r term 90 a" with precedence 19 -for @{ Or_intro_r ?? (cast ? $a ?) }. -notation > "∨\sub\e term 90 ab [ident Ha] term 90 c1 [ident Hb] term 90 c2" with precedence 19 -for @{ Or_elim ??? (cast $ab $ab ?) (λ${ident Ha}.cast ? $c1 ?) (λ${ident Hb}.cast ? $c2 ?) }. -notation > "∀\sub\i [ident z] term 90 a" with precedence 19 -for @{ Forall_intro ?? (λ${ident z}.cast ? $a ?) }. -notation > "∀\sub\e term 90 ab" with precedence 19 -for @{ Forall_elim ?? ? (cast $ab $ab ?) }. (* CSC: WRONG *) -notation > "∃\sub\e term 90 enpn [ident z] [ident pz] term 90 c" with precedence 19 -for @{ Exists_elim ??? (cast $enpn $enpn ?) (λ${ident z}.λ${ident pz}.cast ? $c ?) }. -notation > "∃\sub\i term 90 n term 90 pn" with precedence 19 -for @{ Exists_intro ? (λ_.?) $n (cast ? $pn ?) }. - -lemma ex1 : (A ⇒ E) ∨ B ⇒ A ∧ C ⇒ (E ∧ C) ∨ B. - apply (cast ? ((A⇒E)∨B⇒A∧C⇒E∧C∨B)); - (*NICE: TRY THIS ERROR! - apply (⇒\sub\i [H] (A∧C⇒E∧E∧C∨B)); - apply (⇒\sub\i [K] (E∧E∧C∨B)); + (*NICE: TRY THIS ERROR! + apply (⇒_i [H] (A∧C⇒E∧E∧C∨B)); + apply (⇒_i [K] (E∧E∧C∨B)); OR DO THE RIGHT THING *) - apply (⇒\sub\i [H] (A∧C⇒E∧C∨B)); - apply (⇒\sub\i [K] (E∧C∨B)); - - apply (∨\sub\e ((A⇒E)∨B) [C1] (E∧C∨B) [C2] (E∧C∨B)); + apply (⇒_i [H] (A∧C⇒E∧C∨B)); + apply (⇒_i [K] (E∧C∨B)); + apply (∨_e ((A⇒E)∨B) [C1] (E∧C∨B) [C2] (E∧C∨B)); [ apply [H]; -| apply (∨\sub\i\sup\l (E∧C)); - apply (∧\sub\i E C); - [ apply (⇒\sub\e (A⇒E) A); +| apply (∨_i_l (E∧C)); + apply (∧_i E C); + [ apply (⇒_e (A⇒E) A); [ apply [C1]; - | apply (∧\sub\e\sup\l (A∧C)); - apply [K]; + | apply (∧_e_l (A∧C)); apply [K]; ] - | apply (∧\sub\e\sup\r (A∧C)); - apply [K]; + | apply (∧_e_r (A∧C)); apply [K]; ] -| apply (∨\sub\i\sup\r B); - apply [C2]; +| apply (∨_i_r B); apply [C2]; ] qed. -axiom N: Type. -axiom R: N → N → CProp. +lemma ex2: ΠN:Type.ΠR:N→N→CProp. -lemma ex2: (∀a:N.∀b:N.R a b ⇒ R b a) ⇒ ∀z:N.(∃x.R x z) ⇒ ∃y. R z y. - apply (cast ? ((∀a:N.∀b:N.R a b ⇒ R b a) ⇒ ∀z:N.(∃x.R x z) ⇒ ∃y. R z y)); - apply (⇒\sub\i [H] (∀z:N.(∃x:N.R x z)⇒∃y:N.R z y)); - apply (∀\sub\i [z] ((∃x:N.R x z)⇒∃y:N.R z y)); - apply (⇒\sub\i [H2] (∃y:N.R z y)); - apply (∃\sub\e (∃x:N.R x z) [n] [H3] (∃y:N.R z y)); + (∀a:N.∀b:N.R a b ⇒ R b a) ⇒ ∀z:N.(∃x.R x z) ⇒ ∃y. R z y. + + intros (N R);apply (prove ((∀a:N.∀b:N.R a b ⇒ R b a) ⇒ ∀z:N.(∃x.R x z) ⇒ ∃y. R z y)); + + apply (⇒_i [H] (∀z:N.(∃x:N.R x z)⇒∃y:N.R z y)); + apply (∀_i [z] ((∃x:N.R x z)⇒∃y:N.R z y)); + apply (⇒_i [H2] (∃y:N.R z y)); + apply (∃_e (∃x:N.R x z) [n] [H3] (∃y:N.R z y)); [ apply [H2] - | apply (∃\sub\i n (R z n)); - apply (⇒\sub\e (R n z ⇒ R z n) (R n z)); - [ apply (∀\sub\e (∀b:N.R n b ⇒ R b n)); - apply (∀\sub\e (∀a:N.∀b:N.R a b ⇒ R b a)); + | apply (∃_i n (R z n)); + apply (⇒_e (R n z ⇒ R z n) (R n z)); + [ apply (∀_e (∀b:N.R n b ⇒ R b n) z); + apply (∀_e (∀a:N.∀b:N.R a b ⇒ R b a) n); apply [H] | apply [H3] ] diff --git a/helm/software/matita/library/demo/natural_deduction_support.ma b/helm/software/matita/library/demo/natural_deduction_support.ma new file mode 100644 index 000000000..a3613e51d --- /dev/null +++ b/helm/software/matita/library/demo/natural_deduction_support.ma @@ -0,0 +1,158 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +axiom cast: ∀A,B:CProp.B → A. + +notation > "'prove' p" non associative with precedence 19 for @{ cast ? $p}. + +notation < "\infrule (t\atop ⋮) mstyle color #ff0000 (b) (? \ERROR)" with precedence 19 +for @{ 'caste $a $b $t }. + +interpretation "cast ERR" 'caste a b t = (cast a b t). + +notation < "\infrule (t\atop ⋮) a ?" with precedence 19 for @{ 'cast $a $t }. + +interpretation "cast OK" 'cast a t = (cast a a t). + +definition assumpt ≝ λA:CProp.λa:A.a. + +notation < "[ a ] \sup mstyle color #ff0000 (H)" with precedence 19 for @{ 'asse $a $H }. +interpretation "assumption" 'asse a H = (cast _ _ (assumpt a (cast _ _ H))). + +notation < "[ a ] \sup H" with precedence 19 for @{ 'ass $a $H }. +interpretation "assumption" 'ass a H = (cast a a (assumpt a (cast a a H))). + +inductive Imply (A,B:CProp) : CProp ≝ + Imply_intro: (A → B) → Imply A B. + +notation "hbox(a break ⇒ b)" right associative with precedence 20 for @{ 'Imply $a $b }. +interpretation "Imply" 'Imply a b = (Imply a b). + +notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000(ab) (⇒\sub\i \emsp ident H \ERROR) " with precedence 19 +for @{ 'Imply_introe $xxx $ab (λ${ident H}:$p.$b) }. +interpretation "Imply_intro" 'Imply_introe xxx ab \eta.b = (cast xxx ab (Imply_intro _ _ b)). + +notation < "\infrule hbox(\emsp b \emsp) ab (⇒\sub\i \emsp ident H) " with precedence 19 +for @{ 'Imply_intro $ab (λ${ident H}:$p.$b) }. +interpretation "Imply_intro" 'Imply_intro ab \eta.b = (cast ab ab (Imply_intro _ _ b)). + +definition Imply_elim ≝ λA,B.λf:Imply A B.λa:A.match f with [ Imply_intro g ⇒ g a]. + +notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b (⇒\sub\e) " with precedence 19 for @{ 'Imply_elim $ab $a $b }. +interpretation "Imply_elim" 'Imply_elim ab a b = (cast _ b (Imply_elim _ _ ab a)). + +inductive And (A,B:CProp) : CProp ≝ + And_intro: A → B → And A B. + +interpretation "constructive and" 'and x y = (And x y). + +notation < "\infrule hbox(\emsp a \emsp\emsp\emsp b \emsp) ab (∧\sub\i)" with precedence 19 for @{ 'And_intro $a $b $ab }. +interpretation "And_intro" 'And_intro a b ab = (cast _ ab (And_intro _ _ a b)). + +definition And_elim_l ≝ + λA,B.λc:A∧B.match c with [ And_intro a b ⇒ a ]. + +notation < "\infrule hbox(\emsp ab \emsp) a (∧\sub(\e_\l))" with precedence 19 for @{ 'And_elim_l $ab $a }. +interpretation "And_elim_l" 'And_elim_l ab a = (cast _ a (And_elim_l _ _ ab)). + +definition And_elim_r ≝ + λA,B.λc:A∧B.match c with [ And_intro a b ⇒ b ]. + +notation < "\infrule hbox(\emsp ab \emsp) b (∧\sub(\e_\r))" with precedence 19 for @{ 'And_elim_r $ab $b }. +interpretation "And_elim_r" 'And_elim_r ab b = (cast _ b (And_elim_r _ _ ab)). + +inductive Or (A,B:CProp) : CProp ≝ + | Or_intro_l: A → Or A B + | Or_intro_r: B → Or A B. + +interpretation "constructive or" 'or x y = (Or x y). + +notation < "\infrule hbox(\emsp a \emsp) ab (∨\sub(\i_\l))" with precedence 19 for @{ 'Or_intro_l $a $ab }. +interpretation "Or_intro_l" 'Or_intro_l a ab = (cast _ ab (Or_intro_l _ _ a)). + +notation < "\infrule hbox(\emsp b \emsp) ab (∨\sub(\i_\r))" with precedence 19 for @{ 'Or_intro_r $b $ab }. +interpretation "Or_intro_r" 'Or_intro_r b ab = (cast _ ab (Or_intro_r _ _ b)). + +definition Or_elim ≝ + λA,B,C:CProp.λc:A∨B.λfa: A → C.λfb: B → C. + match c with [ Or_intro_l a ⇒ fa a | Or_intro_r b ⇒ fb b]. + +notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp ac \emsp\emsp\emsp bc \emsp) c (∨\sub\e \emsp ident Ha \emsp ident Hb)" with precedence 19 +for @{ 'Or_elim $ab (λ${ident Ha}:$ta.$ac) (λ${ident Hb}:$tb.$bc) $c }. +interpretation "Or_elim" 'Or_elim ab \eta.ac \eta.bc c = (cast _ c (Or_elim _ _ _ ab ac bc)). + +inductive Exists (A:Type) (P:A→CProp) : CProp ≝ + Exists_intro: ∀w:A. P w → Exists A P. + +interpretation "constructive ex" 'exists \eta.x = (Exists _ x). + +notation < "\infrule hbox(\emsp Pn \emsp) Px (∃\sub\i)" with precedence 19 +for @{ 'Exists_intro $Pn $Px }. +interpretation "Exists_intro" 'Exists_intro Pn Px = (cast _ Px (Exists_intro _ _ _ Pn)). + +definition Exists_elim ≝ + λA:Type.λP:A→CProp.λC:CProp.λc:∃x:A.P x.λH:(∀x.P x → C). + match c with [ Exists_intro w p ⇒ H w p ]. + +notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) c (∃\sub\e \emsp ident n \emsp ident HPn)" with precedence 19 +for @{ 'Exists_elim $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }. +interpretation "Exists_elim" 'Exists_elim ExPx Pc c = (cast _ c (Exists_elim _ _ _ ExPx Pc)). + +inductive Forall (A:Type) (P:A→CProp) : CProp ≝ + Forall_intro: (∀n:A. P n) → Forall A P. + +notation "\forall ident x:A.break term 19 Px" with precedence 20 +for @{ 'Forall (λ${ident x}:$A.$Px) }. +interpretation "Forall" 'Forall \eta.Px = (Forall _ Px). + +notation < "\infrule hbox(\emsp Px \emsp) Pn (∀\sub\i \emsp ident x)" with precedence 19 +for @{ 'Forall_intro (λ${ident x}:$tx.$Px) $Pn }. +interpretation "Forall_intro" 'Forall_intro Px Pn = (cast _ Pn (Forall_intro _ _ Px)). + +definition Forall_elim ≝ + λA:Type.λP:A→CProp.λn:A.λf:∀x:A.P x.match f with [ Forall_intro g ⇒ g n ]. + +notation < "\infrule hbox(\emsp Px \emsp) Pn (∀\sub\i)" with precedence 19 for @{ 'Forall_elim $Px $Pn }. +interpretation "Forall_elim" 'Forall_elim Px Pn = (cast _ Pn (Forall_elim _ _ _ Px)). + +notation > "[H]" with precedence 90 +for @{ assumpt ? (cast ? ? $H)}. +notation > "⇒_'i' [ident H] term 90 b" with precedence 19 +for @{ Imply_intro ?? (λ${ident H}.cast ? $b ?) }. +notation > "⇒_'e' term 90 ab term 90 a" with precedence 19 +for @{ Imply_elim ?? (cast ? $ab ?) (cast $a $a ?) }. +notation > "∧_'i' term 90 a term 90 b" with precedence 19 +for @{ And_intro ?? (cast ? $a ?) (cast ? $b ?) }. +(*notation > "∧_'e_l' term 90 ab" with precedence 19 +for @{ And_elim_l ?? (cast (? ∧ False) $ab ?) }. +notation > "∧_'e_l' term 90 a ∧ term 90 b" with precedence 19 +for @{ And_elim_l ?? (cast (? ∧ $b) ($a ∧ $b) ?) }. *) +notation > "∧_'e_l' term 90 ab" with precedence 19 +for @{ And_elim_l ?? (cast $ab $ab ?) }. (* CSC: WRONG *) +notation > "∧_'e_r' term 90 ab" with precedence 19 +for @{ And_elim_r ?? (cast $ab $ab ?) }. (* CSC: WRONG *) +notation > "∨_'i_l' term 90 a" with precedence 19 +for @{ Or_intro_l ?? (cast ? $a ?) }. +notation > "∨_'i_r' term 90 a" with precedence 19 +for @{ Or_intro_r ?? (cast ? $a ?) }. +notation > "∨_'e' term 90 ab [ident Ha] term 90 c1 [ident Hb] term 90 c2" with precedence 19 +for @{ Or_elim ??? (cast $ab $ab ?) (λ${ident Ha}.cast ? $c1 ?) (λ${ident Hb}.cast ? $c2 ?) }. +notation > "∀_'i' [ident z] term 90 a" with precedence 19 +for @{ Forall_intro ?? (λ${ident z}.cast ? $a ?) }. +notation > "∀_'e' term 90 ab term 90 a" with precedence 19 +for @{ Forall_elim ?? $a (cast $ab $ab ?) }. +notation > "∃_'e' term 90 enpn [ident z] [ident pz] term 90 c" with precedence 19 +for @{ Exists_elim ??? (cast $enpn $enpn ?) (λ${ident z}.λ${ident pz}.cast ? $c ?) }. +notation > "∃_'i' term 90 n term 90 pn" with precedence 19 +for @{ Exists_intro ? (λ_.?) $n (cast ? $pn ?) }. diff --git a/helm/software/matita/library/depends b/helm/software/matita/library/depends index ae3231d91..27032d916 100644 --- a/helm/software/matita/library/depends +++ b/helm/software/matita/library/depends @@ -1,109 +1,111 @@ -Z/times.ma Z/plus.ma nat/lt_arith.ma -list/sort.ma datatypes/bool.ma datatypes/constructors.ma list/in.ma -list/list.ma datatypes/bool.ma higher_order_defs/functions.ma logic/equality.ma nat/nat.ma nat/orders.ma nat/plus.ma -nat/factorial.ma nat/le_arith.ma -nat/sigma_and_pi.ma nat/exp.ma nat/factorial.ma nat/lt_arith.ma -nat/compare.ma datatypes/bool.ma datatypes/compare.ma nat/orders.ma -nat/neper.ma nat/binomial.ma nat/chebyshev.ma nat/div_and_mod_diseq.ma nat/iteration2.ma nat/log.ma +formal_topology/formal_topologies.ma formal_topology/basic_topologies.ma +demo/formal_topology.ma logic/cprop_connectives.ma logic/equality.ma Q/ratio/rtimes.ma Q/fraction/ftimes.ma Q/ratio/rinv.ma +demo/power_derivative.ma nat/compare.ma nat/nat.ma nat/orders.ma nat/plus.ma +nat/compare.ma datatypes/bool.ma datatypes/compare.ma nat/orders.ma nat/lt_arith.ma nat/div_and_mod.ma +demo/propositional_sequent_calculus.ma datatypes/constructors.ma list/sort.ma nat/compare.ma nat/plus.ma +Z/inversion.ma Z/dirichlet_product.ma Z/moebius.ma +nat/factorial2.ma nat/exp.ma nat/factorial.ma +nat/orders.ma higher_order_defs/ordering.ma logic/connectives.ma nat/nat.ma +nat/sieve.ma list/sort.ma nat/primes.ma +technicalities/setoids.ma datatypes/constructors.ma logic/coimplication.ma logic/connectives2.ma +nat/div_and_mod_diseq.ma nat/lt_arith.ma +logic/cprop_connectives.ma logic/connectives.ma +algebra/groups.ma algebra/monoids.ma datatypes/bool.ma logic/connectives.ma nat/compare.ma nat/le_arith.ma +nat/chinese_reminder.ma nat/congruence.ma nat/exp.ma nat/gcd.ma nat/permutation.ma +formal_topology/saturations_reductions.ma datatypes/subsets.ma +Q/q/qinv.ma Q/q/q.ma Q/ratio/rinv.ma nat/exp.ma nat/div_and_mod.ma nat/lt_arith.ma -nat/binomial.ma nat/factorial2.ma nat/iteration2.ma -Q/q/qtimes.ma Q/q/qinv.ma Q/ratio/rtimes.ma -nat/euler_theorem.ma nat/map_iter_p.ma nat/nat.ma nat/totient.ma -Z/orders.ma Z/z.ma nat/orders.ma -datatypes/subsets.ma datatypes/categories.ma logic/cprop_connectives.ma -logic/connectives2.ma higher_order_defs/relations.ma +list/in.ma datatypes/bool.ma datatypes/constructors.ma list/list.ma +datatypes/compare.ma nat/factorization2.ma list/list.ma nat/factorization.ma nat/sieve.ma -Q/ratio/rinv.ma Q/fraction/finv.ma Q/ratio/ratio.ma -nat/o.ma nat/binomial.ma nat/sqrt.ma -nat/permutation.ma nat/compare.ma nat/sigma_and_pi.ma -Q/q/qinv.ma Q/q/q.ma Q/ratio/rinv.ma -nat/plus.ma nat/nat.ma -Q/fraction/fraction.ma Z/compare.ma nat/factorization.ma -formal_topology/saturations_reductions.ma datatypes/subsets.ma -algebra/finite_groups.ma algebra/groups.ma nat/relevant_equations.ma -Z/dirichlet_product.ma Z/sigma_p.ma Z/times.ma nat/primes.ma -datatypes/categories.ma logic/cprop_connectives.ma -nat/chinese_reminder.ma nat/congruence.ma nat/exp.ma nat/gcd.ma nat/permutation.ma -nat/fermat_little_theorem.ma nat/congruence.ma nat/exp.ma nat/gcd.ma nat/permutation.ma -nat/div_and_mod.ma datatypes/constructors.ma nat/minus.ma -nat/count.ma nat/permutation.ma nat/relevant_equations.ma nat/sigma_and_pi.ma -higher_order_defs/ordering.ma logic/equality.ma -datatypes/constructors.ma logic/equality.ma -Q/q.ma Q/fraction/fraction.ma Z/compare.ma Z/plus.ma nat/factorization.ma -Z/compare.ma Z/orders.ma nat/compare.ma -Q/q/q.ma Q/fraction/numerator_denominator.ma Q/ratio/ratio.ma -Q/fraction/numerator_denominator.ma Q/fraction/finv.ma -nat/log.ma datatypes/constructors.ma nat/div_and_mod_diseq.ma nat/iteration2.ma nat/minimization.ma nat/primes.ma nat/relevant_equations.ma -datatypes/bool.ma higher_order_defs/functions.ma logic/equality.ma -nat/chebyshev_teta.ma nat/binomial.ma nat/pi_p.ma -Q/ratio/ratio.ma Q/fraction/fraction.ma -nat/div_and_mod_diseq.ma nat/lt_arith.ma -decidable_kit/list_aux.ma decidable_kit/eqtype.ma list/list.ma nat/plus.ma +formal_topology/basic_topologies.ma datatypes/categories.ma formal_topology/relations.ma formal_topology/saturations_reductions.ma logic/connectives.ma -technicalities/setoids.ma datatypes/constructors.ma logic/coimplication.ma logic/connectives2.ma -datatypes/compare.ma -nat/orders.ma higher_order_defs/ordering.ma logic/connectives.ma nat/nat.ma -nat/sqrt.ma nat/compare.ma nat/log.ma nat/times.ma -nat/relevant_equations.ma nat/gcd.ma nat/minus.ma nat/times.ma -nat/factorial2.ma nat/exp.ma nat/factorial.ma +Q/nat_fact/times.ma nat/factorization.ma +decidable_kit/fintype.ma decidable_kit/eqtype.ma decidable_kit/list_aux.ma +demo/natural_deduction_support.ma +nat/ord.ma datatypes/constructors.ma nat/exp.ma nat/gcd.ma nat/nth_prime.ma nat/relevant_equations.ma +nat/totient1.ma nat/compare.ma nat/gcd_properties1.ma nat/iteration2.ma nat/totient.ma +nat/times.ma nat/plus.ma nat/chebyshev_thm.ma nat/neper.ma -algebra/groups.ma algebra/monoids.ma datatypes/bool.ma logic/connectives.ma nat/compare.ma nat/le_arith.ma -demo/power_derivative.ma nat/compare.ma nat/nat.ma nat/orders.ma nat/plus.ma -nat/gcd_properties1.ma nat/gcd.ma -nat/minimization.ma nat/minus.ma +Z/z.ma datatypes/bool.ma nat/nat.ma +demo/cantor.ma datatypes/constructors.ma demo/formal_topology.ma +nat/nth_prime.ma nat/lt_arith.ma nat/primes.ma +nat/le_arith.ma nat/orders.ma nat/times.ma decidable_kit/fgraph.ma decidable_kit/fintype.ma -formal_topology/relations.ma datatypes/subsets.ma -formal_topology/formal_topologies.ma formal_topology/basic_topologies.ma -Q/fraction/finv.ma Q/fraction/fraction.ma Z/plus.ma +nat/euler_theorem.ma nat/map_iter_p.ma nat/nat.ma nat/totient.ma +Q/fraction/ftimes.ma Q/fraction/finv.ma Q/nat_fact/times.ma Q/ratio/ratio.ma Z/times.ma +nat/factorial.ma nat/le_arith.ma Z/plus.ma Z/z.ma nat/minus.ma -nat/minus.ma nat/compare.ma nat/le_arith.ma -nat/generic_iter_p.ma nat/div_and_mod_diseq.ma nat/ord.ma nat/primes.ma -Q/Qaxioms.ma Z/compare.ma Z/times.ma nat/iteration2.ma -Z/z.ma datatypes/bool.ma nat/nat.ma -formal_topology/concrete_spaces.ma formal_topology/basic_pairs.ma +Q/ratio/rinv.ma Q/fraction/finv.ma Q/ratio/ratio.ma decidable_kit/streicher.ma logic/connectives.ma logic/equality.ma -nat/ord.ma datatypes/constructors.ma nat/exp.ma nat/gcd.ma nat/nth_prime.ma nat/relevant_equations.ma +nat/fermat_little_theorem.ma nat/congruence.ma nat/exp.ma nat/gcd.ma nat/permutation.ma +Q/q/qplus.ma nat/factorization.ma +decidable_kit/list_aux.ma decidable_kit/eqtype.ma list/list.ma nat/plus.ma nat/map_iter_p.ma nat/count.ma nat/permutation.ma +Q/q.ma Q/fraction/fraction.ma Z/compare.ma Z/plus.ma nat/factorization.ma +Z/orders.ma Z/z.ma nat/orders.ma +nat/permutation.ma nat/compare.ma nat/sigma_and_pi.ma demo/realisability.ma datatypes/constructors.ma logic/connectives.ma -nat/totient1.ma nat/compare.ma nat/gcd_properties1.ma nat/iteration2.ma nat/totient.ma -decidable_kit/eqtype.ma datatypes/constructors.ma decidable_kit/decidable.ma -nat/gcd.ma nat/lt_arith.ma nat/primes.ma -nat/factorization.ma nat/ord.ma -algebra/monoids.ma algebra/semigroups.ma -demo/propositional_sequent_calculus.ma datatypes/constructors.ma list/sort.ma nat/compare.ma nat/plus.ma -logic/coimplication.ma logic/connectives.ma -decidable_kit/fintype.ma decidable_kit/eqtype.ma decidable_kit/list_aux.ma -logic/equality.ma higher_order_defs/relations.ma logic/connectives.ma -higher_order_defs/functions.ma logic/equality.ma -algebra/semigroups.ma higher_order_defs/functions.ma -nat/sieve.ma list/sort.ma nat/primes.ma -nat/pi_p.ma nat/generic_iter_p.ma nat/iteration2.ma nat/primes.ma -Q/nat_fact/times.ma nat/factorization.ma +list/list.ma datatypes/bool.ma higher_order_defs/functions.ma logic/equality.ma nat/nat.ma nat/orders.ma nat/plus.ma nat/totient.ma nat/chinese_reminder.ma nat/iteration2.ma -logic/cprop_connectives.ma logic/connectives.ma -nat/primes.ma logic/connectives.ma nat/div_and_mod.ma nat/factorial.ma nat/minimization.ma nat/sigma_and_pi.ma -higher_order_defs/relations.ma logic/connectives.ma -demo/toolbox.ma logic/cprop_connectives.ma -demo/natural_deduction.ma -Q/inv.ma Q/q/q.ma Q/fraction/finv.ma Q/q.ma +nat/sigma_and_pi.ma nat/exp.ma nat/factorial.ma nat/lt_arith.ma +nat/count.ma nat/permutation.ma nat/relevant_equations.ma nat/sigma_and_pi.ma +formal_topology/basic_pairs.ma datatypes/categories.ma formal_topology/relations.ma +Q/frac.ma Q/q/q.ma Q/q/qinv.ma nat/factorization.ma nat/nat.ma +Q/q/qtimes.ma Q/q/qinv.ma Q/ratio/rtimes.ma +formal_topology/concrete_spaces.ma formal_topology/basic_pairs.ma +nat/minus.ma nat/compare.ma nat/le_arith.ma +Q/ratio/ratio.ma Q/fraction/fraction.ma +nat/chebyshev_teta.ma nat/binomial.ma nat/pi_p.ma +algebra/finite_groups.ma algebra/groups.ma nat/relevant_equations.ma decidable_kit/decidable.ma datatypes/bool.ma decidable_kit/streicher.ma logic/connectives.ma nat/compare.ma +nat/pi_p.ma nat/generic_iter_p.ma nat/iteration2.ma nat/primes.ma +algebra/semigroups.ma higher_order_defs/functions.ma +higher_order_defs/relations.ma logic/connectives.ma +nat/factorization.ma nat/ord.ma +formal_topology/relations.ma datatypes/subsets.ma +nat/neper.ma nat/binomial.ma nat/chebyshev.ma nat/div_and_mod_diseq.ma nat/iteration2.ma nat/log.ma Z/moebius.ma Z/sigma_p.ma nat/factorization.ma -Z/inversion.ma Z/dirichlet_product.ma Z/moebius.ma -Q/fraction/ftimes.ma Q/fraction/finv.ma Q/nat_fact/times.ma Q/ratio/ratio.ma Z/times.ma -Q/q/qplus.ma nat/factorization.ma -Q/frac.ma Q/q/q.ma Q/q/qinv.ma nat/nat.ma nat/factorization.ma -Z/sigma_p.ma Z/plus.ma Z/times.ma nat/generic_iter_p.ma nat/ord.ma nat/primes.ma +demo/toolbox.ma logic/cprop_connectives.ma +nat/iteration2.ma nat/count.ma nat/generic_iter_p.ma nat/ord.ma nat/primes.ma +logic/coimplication.ma logic/connectives.ma +nat/minimization.ma nat/minus.ma +logic/connectives2.ma higher_order_defs/relations.ma +datatypes/subsets.ma datatypes/categories.ma logic/cprop_connectives.ma +nat/chebyshev.ma nat/factorial2.ma nat/factorization.ma nat/log.ma nat/o.ma nat/pi_p.ma +decidable_kit/eqtype.ma datatypes/constructors.ma decidable_kit/decidable.ma +Q/q/q.ma Q/fraction/numerator_denominator.ma Q/ratio/ratio.ma nat/bertrand.ma list/in.ma list/sort.ma nat/chebyshev.ma nat/chebyshev_teta.ma nat/o.ma nat/sieve.ma nat/sqrt.ma nat/nat.ma higher_order_defs/functions.ma -formal_topology/basic_topologies.ma datatypes/categories.ma formal_topology/relations.ma formal_topology/saturations_reductions.ma -nat/chebyshev.ma nat/o.ma nat/factorial2.ma nat/factorization.ma nat/log.ma nat/pi_p.ma -demo/formal_topology.ma logic/cprop_connectives.ma logic/equality.ma +Q/Qaxioms.ma Z/compare.ma Z/times.ma nat/iteration2.ma +demo/natural_deduction.ma +higher_order_defs/ordering.ma logic/equality.ma nat/congruence.ma nat/primes.ma nat/relevant_equations.ma -nat/le_arith.ma nat/orders.ma nat/times.ma -nat/times.ma nat/plus.ma -formal_topology/basic_pairs.ma datatypes/categories.ma formal_topology/relations.ma -list/in.ma datatypes/bool.ma datatypes/constructors.ma list/list.ma -nat/nth_prime.ma nat/lt_arith.ma nat/primes.ma -nat/iteration2.ma nat/count.ma nat/generic_iter_p.ma nat/ord.ma nat/primes.ma +logic/equality.ma logic/connectives.ma higher_order_defs/relations.ma +nat/gcd.ma nat/lt_arith.ma nat/primes.ma +datatypes/bool.ma higher_order_defs/functions.ma logic/equality.ma +Z/compare.ma Z/orders.ma nat/compare.ma +algebra/monoids.ma algebra/semigroups.ma +nat/div_and_mod.ma datatypes/constructors.ma nat/minus.ma +Z/dirichlet_product.ma Z/sigma_p.ma Z/times.ma nat/primes.ma +nat/sqrt.ma nat/compare.ma nat/log.ma nat/times.ma +datatypes/categories.ma logic/cprop_connectives.ma +nat/relevant_equations.ma nat/gcd.ma nat/minus.ma nat/times.ma +Q/fraction/finv.ma Q/fraction/fraction.ma Z/plus.ma +nat/binomial.ma nat/factorial2.ma nat/iteration2.ma +nat/log.ma datatypes/constructors.ma nat/div_and_mod_diseq.ma nat/iteration2.ma nat/minimization.ma nat/primes.ma nat/relevant_equations.ma +higher_order_defs/functions.ma logic/equality.ma +Q/fraction/numerator_denominator.ma Q/fraction/finv.ma +nat/generic_iter_p.ma nat/div_and_mod_diseq.ma nat/ord.ma nat/primes.ma +datatypes/constructors.ma logic/equality.ma +nat/plus.ma nat/nat.ma +Q/fraction/fraction.ma Z/compare.ma nat/factorization.ma +nat/primes.ma logic/connectives.ma nat/div_and_mod.ma nat/factorial.ma nat/minimization.ma nat/sigma_and_pi.ma +nat/gcd_properties1.ma nat/gcd.ma +list/sort.ma datatypes/bool.ma datatypes/constructors.ma list/in.ma +Z/times.ma Z/plus.ma nat/lt_arith.ma +Z/sigma_p.ma Z/plus.ma Z/times.ma nat/generic_iter_p.ma nat/ord.ma nat/primes.ma +nat/o.ma nat/binomial.ma nat/sqrt.ma +Q/inv.ma Q/fraction/finv.ma Q/q/q.ma Q/q.ma -- 2.39.2