]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/didactic/support/natural_deduction.ma
print an error message if graphviz is not found
[helm.git] / helm / software / matita / library / didactic / support / natural_deduction.ma
index 3e7c54d5cbbd6873921b72b9fc6060a6f774ac43..84fbeee781a535621476e798224d1c662cce5a81 100644 (file)
@@ -45,6 +45,21 @@ definition Discharge := λA:CProp.λa:A.
 
 axiom Raa : ∀A.(Not A → Bot) → A.
 
+axiom sort : Type.
+
+inductive Exists (A:Type) (P:A→CProp) : CProp ≝
+  Exists_intro: ∀w:A. P w → Exists A P.
+
+definition Exists_elim ≝
+  λA:Type.λP:A→CProp.λC:CProp.λc:Exists A P.λH:(Πx.P x → C).
+   match c with [ Exists_intro w p ⇒ H w p ].
+
+inductive Forall (A:Type) (P:A→CProp) : CProp ≝
+ Forall_intro: (∀n:A. P n) → Forall A P.
+
+definition Forall_elim ≝
+ λA:Type.λP:A→CProp.λn:A.λf:Forall A P.match f with [ Forall_intro g ⇒ g n ].
+
 (* Dummy proposition *)
 axiom unit : CProp.
 
@@ -61,7 +76,17 @@ interpretation "Bot" 'Bot = Bot.
 interpretation "Not" 'not a = (Not a).
 notation "✶" non associative with precedence 90 for @{'unit}.
 interpretation "dummy prop" 'unit = unit.
-
+notation > "\exists list1 ident x sep , . term 19 Px" with precedence 20
+for ${ fold right @{$Px} rec acc @{'myexists (λ${ident x}.$acc)} }.
+notation < "hvbox(\exists ident i break . p)" with precedence 20
+for @{ 'myexists (\lambda ${ident i} : $ty. $p) }.
+interpretation "constructive ex" 'myexists \eta.x = (Exists sort x).
+notation > "\forall ident x.break term 19 Px" with precedence 20
+for @{ 'Forall (λ${ident x}.$Px) }.
+notation < "\forall ident x.break term 19 Px" with precedence 20
+for @{ 'Forall (λ${ident x}:$tx.$Px) }.
+interpretation "Forall" 'Forall \eta.Px = (Forall _ Px).
 (* Variables *)
 axiom A : CProp.
 axiom B : CProp.
@@ -78,24 +103,20 @@ axiom L : CProp.
 axiom M : CProp.
 axiom N : CProp.
 axiom O : CProp.
-axiom P : CProp.
-axiom Q : CProp.
-axiom R : CProp.
-axiom S : CProp.
-axiom T : CProp.
-axiom U : CProp.
-axiom V : CProp.
-axiom W : CProp.
-axiom X : CProp.
-axiom Y : CProp.
-axiom Z : CProp.
+axiom P : sort →CProp.
+axiom Q : sort →CProp.
+axiom R : sort →sort →CProp.
+axiom S : sort →sort →CProp.
+axiom f : sort → sort.
+axiom g : sort → sort.
+axiom h : sort → sort → sort.
 
 (* Every formula user provided annotates its proof:
    `A` becomes `(show A ?)` *)
-definition show : A.A→A ≝ λA:CProp.λa:A.a.
+definition show : ΠA.A→A ≝ λA:CProp.λa:A.a.
 
 (* When something does not fit, this daemon is used *)
-axiom cast: A,B:CProp.B → A.
+axiom cast: ΠA,B:CProp.B → A.
 
 (* begin a proof: draws the root *)
 notation > "'prove' p" non associative with precedence 19 
@@ -111,6 +132,43 @@ notation < "\infrule (t\atop ⋮) mstyle color #ff0000 (a) ?" with precedence 19
 for @{ 'leaf_ko $a $t }.
 interpretation "leaf KO" 'leaf_ko a t = (cast _ _ (show a t)).
 
+(* already proved lemma *)
+definition Lemma : ΠA.A→A ≝ λA:CProp.λa:A.a.
+
+notation < "\infrule \nbsp p mstyle color #ff0000 (\lem\emsp H)" non associative with precedence 19
+for @{ 'lemma_ko_1 $p $H }.
+interpretation "lemma_ko_1" 'lemma_ko_1 p H = 
+  (show p (cast _ _ (Lemma _ H))). 
+
+notation < "\infrule \nbsp mstyle color #ff0000 (p) mstyle color #ff0000 (\lem\emsp H)" non associative with precedence 19
+for @{ 'lemma_ko_2 $p $H }.
+interpretation "lemma_ko_2" 'lemma_ko_2 p H = 
+  (cast _ _ (show p (cast _ _ (Lemma _ H)))). 
+
+notation < "\infrule \nbsp p (\lem\emsp H)" non associative with precedence 19
+for @{ 'lemma_ok_1 $p $H }.
+interpretation "lemma_ok_1" 'lemma_ok_1 p H = 
+  (show p (Lemma _ H)). 
+interpretation "lemma_ok_11" 'lemma_ok_1 p H = 
+  (show p (Lemma _ H _)). 
+interpretation "lemma_ok_111" 'lemma_ok_1 p H = 
+  (show p (Lemma _ H _ _)). 
+
+notation < "\infrule \nbsp mstyle color #ff0000 (p) (\lem\emsp H)" non associative with precedence 19
+for @{ 'lemma_ok_2 $p $H }.
+interpretation "lemma_ok_2" 'lemma_ok_2 p H = 
+  (cast _ _ (show p (Lemma _ H))). 
+interpretation "lemma_ok_22" 'lemma_ok_2 p H = 
+  (cast _ _ (show p (Lemma _ H _))). 
+interpretation "lemma_ok_22" 'lemma_ok_2 p H = 
+  (cast _ _ (show p (Lemma _ H _ _))). 
+
+notation > "'lem' term 90 p" non associative with precedence 19
+for @{ 'Lemma $p  }.
+interpretation "lemma KO" 'Lemma p = (cast _ _ (Lemma _ p)).
+interpretation "lemma OK" 'Lemma p = (Lemma _ p).
+
+
 (* discharging *)
 notation < "[ a ] \sup mstyle color #ff0000 (H)" with precedence 19 
 for @{ 'discharge_ko_1 $a $H }.
@@ -285,7 +343,7 @@ interpretation "Or_intro_l_ok_1" 'Or_intro_l_ok_1 a ab =
   (show ab (Or_intro_l _ _ a)).
 
 notation < "\infrule hbox(\emsp a \emsp) mstyle color #ff0000 (ab) (∨\sub(\i_\l))" with precedence 19 
-for @{ 'Or_intro_l_ok_1 $a $ab }.
+for @{ 'Or_intro_l_ok_2 $a $ab }.
 interpretation "Or_intro_l_ok_2" 'Or_intro_l_ok_2 a ab = 
   (cast _ _ (show ab (Or_intro_l _ _ a))).
 
@@ -310,7 +368,7 @@ interpretation "Or_intro_r_ok_1" 'Or_intro_r_ok_1 a ab =
   (show ab (Or_intro_r _ _ a)).
 
 notation < "\infrule hbox(\emsp a \emsp) mstyle color #ff0000 (ab) (∨\sub(\i_\r))" with precedence 19 
-for @{ 'Or_intro_r_ok_1 $a $ab }.
+for @{ 'Or_intro_r_ok_2 $a $ab }.
 interpretation "Or_intro_r_ok_2" 'Or_intro_r_ok_2 a ab = 
   (cast _ _ (show ab (Or_intro_r _ _ a))).
 
@@ -401,22 +459,22 @@ interpretation "Bot_elim KO" 'Bot_elim a = (Bot_elim _ (cast _ _ a)).
 interpretation "Bot_elim OK" 'Bot_elim a = (Bot_elim _ a).
 
 (* ¬ introduction *)
-notation < "\infrule hbox(\emsp b \emsp) ab (mstyle color #ff0000 (\lnot\sub\i) \emsp ident H)" with precedence 19
+notation < "\infrule hbox(\emsp b \emsp) ab (mstyle color #ff0000 (\lnot\sub(\emsp\i)) \emsp ident H)" with precedence 19
 for @{ 'Not_intro_ko_1 $ab (λ${ident H}:$p.$b) }.
 interpretation "Not_intro_ko_1" 'Not_intro_ko_1 ab \eta.b = 
   (show ab (cast _ _ (Not_intro _ (cast _ _ b)))).
 
-notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (mstyle color #ff0000 (\lnot\sub\i) \emsp ident H)" with precedence 19
+notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (mstyle color #ff0000 (\lnot\sub(\emsp\i)) \emsp ident H)" with precedence 19
 for @{ 'Not_intro_ko_2 $ab (λ${ident H}:$p.$b) }.
 interpretation "Not_intro_ko_2" 'Not_intro_ko_2 ab \eta.b = 
   (cast _ _ (show ab (cast _ _ (Not_intro _ (cast _ _ b))))).
 
-notation < "\infrule hbox(\emsp b \emsp) ab (\lnot\sub\i \emsp ident H) " with precedence 19
+notation < "\infrule hbox(\emsp b \emsp) ab (\lnot\sub(\emsp\i) \emsp ident H) " with precedence 19
 for @{ 'Not_intro_ok_1 $ab (λ${ident H}:$p.$b) }.
 interpretation "Not_intro_ok_1" 'Not_intro_ok_1 ab \eta.b = 
   (show ab (Not_intro _ b)).
 
-notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (\lnot\sub\i \emsp ident H) " with precedence 19
+notation < "\infrule hbox(\emsp b \emsp) mstyle color #ff0000 (ab) (\lnot\sub(\emsp\i) \emsp ident H) " with precedence 19
 for @{ 'Not_intro_ok_2 $ab (λ${ident H}:$p.$b) }.
 interpretation "Not_intro_ok_2" 'Not_intro_ok_2 ab \eta.b = 
   (cast _ _ (show ab (Not_intro _ b))).
@@ -427,22 +485,22 @@ interpretation "Not_intro KO" 'Not_intro a = (cast _ _ (Not_intro _ (cast _ _ a)
 interpretation "Not_intro OK" 'Not_intro a = (Not_intro _ a).
 
 (* ¬ elimination *)
-notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b mstyle color #ff0000 (\lnot\sub\e) " with precedence 19 
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b mstyle color #ff0000 (\lnot\sub(\emsp\e)) " with precedence 19 
 for @{ 'Not_elim_ko_1 $ab $a $b }.
 interpretation "Not_elim_ko_1" 'Not_elim_ko_1 ab a b = 
   (show b (cast _ _ (Not_elim _ (cast _ _ ab) (cast _ _ a)))).
 
-notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) mstyle color #ff0000 (\lnot\sub\e) " with precedence 19 
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) mstyle color #ff0000 (\lnot\sub(\emsp\e)) " with precedence 19 
 for @{ 'Not_elim_ko_2 $ab $a $b }.
 interpretation "Not_elim_ko_2" 'Not_elim_ko_2 ab a b = 
   (cast _ _ (show b (cast _ _ (Not_elim _ (cast _ _ ab) (cast _ _ a))))).
 
-notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b (\lnot\sub\e) " with precedence 19 
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) b (\lnot\sub(\emsp\e)) " with precedence 19 
 for @{ 'Not_elim_ok_1 $ab $a $b }.
 interpretation "Not_elim_ok_1" 'Not_elim_ok_1 ab a b = 
   (show b (Not_elim _ ab a)).
 
-notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) (\lnot\sub\e) " with precedence 19 
+notation < "\infrule hbox(\emsp ab \emsp\emsp\emsp a\emsp) mstyle color #ff0000 (b) (\lnot\sub(\emsp\e)) " with precedence 19 
 for @{ 'Not_elim_ok_2 $ab $a $b }.
 interpretation "Not_elim_ok_2" 'Not_elim_ok_2 ab a b = 
   (cast _ _ (show b (Not_elim _ ab a))).
@@ -450,7 +508,7 @@ interpretation "Not_elim_ok_2" 'Not_elim_ok_2 ab a b =
 notation > "¬_'e' term 90 ab term 90 a" with precedence 19
 for @{ 'Not_elim (show $ab ?) (show $a ?) }.
 interpretation "Not_elim KO" 'Not_elim ab a = 
-  (cast _ _ (Not_elim _ (cast _ _ ab) (cast _ _ a))).
+  (cast _ _ (Not_elim unit (cast _ _ ab) (cast _ _ a))).
 interpretation "Not_elim OK" 'Not_elim ab a = 
   (Not_elim _ ab a).
 
@@ -480,54 +538,119 @@ for @{ 'Raa (λ${ident H}.show $b ?) }.
 interpretation "RAA KO" 'Raa p = (cast _ unit (Raa _ (cast _ (unit_to _) p))).
 interpretation "RAA OK" 'Raa p = (Raa _ p).
 
-(*DOCBEGIN
-Templates for rules:
-
-        apply rule (⇒_i […] (…)).
-
-        apply rule (∧_i (…) (…));
-               [
-               |
-               ]
-
-        apply rule (∨_i_l (…)).
-
-        apply rule (∨_i_r (…)).
-
-        apply rule (¬_i […] (…)).
-
-        apply rule (⊤_i).
-
-        apply rule (⇒_e (…) (…));
-               [
-               |
-               ]
-
-        apply rule (∧_e_l (…)).
-
-        apply rule (∧_e_r (…)).
-
-        apply rule (∨_e (…) […] (…) […] (…));
-               [
-               |
-               |
-               ]
-
-        apply rule (¬_e (…) (…));
-               [
-               |
-               ]
-
-        apply rule (⊥_e (…)).
-
-        apply rule (prove (…)).
-
-        apply rule (RAA […] (…)).
-
-        apply rule (discharge […]).
-
-DOCEND*)
-
-
-
-
+(* ∃ introduction *)
+notation < "\infrule hbox(\emsp Pn \emsp) Px mstyle color #ff0000 (∃\sub\i)" with precedence 19
+for @{ 'Exists_intro_ko_1 $Pn $Px }.
+interpretation "Exists_intro_ko_1" 'Exists_intro_ko_1 Pn Px = 
+  (show Px (cast _ _ (Exists_intro _ _ _ (cast _ _ Pn)))).
+
+notation < "\infrule hbox(\emsp Pn \emsp) mstyle color #ff0000 (Px) mstyle color #ff0000 (∃\sub\i)" with precedence 19
+for @{ 'Exists_intro_ko_2 $Pn $Px }.
+interpretation "Exists_intro_ko_2" 'Exists_intro_ko_2 Pn Px = 
+  (cast _ _ (show Px (cast _ _ (Exists_intro _ _ _ (cast _ _ Pn))))).
+
+notation < "\infrule hbox(\emsp Pn \emsp) Px (∃\sub\i)" with precedence 19
+for @{ 'Exists_intro_ok_1 $Pn $Px }.
+interpretation "Exists_intro_ok_1" 'Exists_intro_ok_1 Pn Px = 
+  (show Px (Exists_intro _ _ _ Pn)).
+
+notation < "\infrule hbox(\emsp Pn \emsp) mstyle color #ff0000 (Px) (∃\sub\i)" with precedence 19
+for @{ 'Exists_intro_ok_2 $Pn $Px }.
+interpretation "Exists_intro_ok_2" 'Exists_intro_ok_2 Pn Px = 
+  (cast _ _ (show Px (Exists_intro _ _ _ Pn))).
+
+notation >"∃_'i' {term 90 t} term 90 Pt" non associative with precedence 19
+for @{'Exists_intro $t (λ_.?) (show $Pt ?)}. 
+interpretation "Exists_intro KO" 'Exists_intro t P Pt =
+ (cast _ _ (Exists_intro sort P t (cast _ _ Pt))).
+interpretation "Exists_intro OK" 'Exists_intro t P Pt =
+ (Exists_intro sort P t Pt).
+(* ∃ elimination *)
+notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) c (mstyle color #ff0000 (∃\sub\e) \emsp ident n \emsp ident HPn)" with precedence 19
+for @{ 'Exists_elim_ko_1 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
+interpretation "Exists_elim_ko_1" 'Exists_elim_ko_1 ExPx Pc c =
+    (show c (cast _ _ (Exists_elim _ _ _ (cast _ _ ExPx) (cast _ _ Pc)))).
+
+notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) mstyle color #ff0000 (c) (mstyle color #ff0000 (∃\sub\e) \emsp ident n \emsp ident HPn)" with precedence 19
+for @{ 'Exists_elim_ko_2 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
+interpretation "Exists_elim_ko_2" 'Exists_elim_ko_2 ExPx Pc c =
+    (cast _ _ (show c (cast _ _ (Exists_elim _ _ _ (cast _ _ ExPx) (cast _ _ Pc))))).
+
+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_ok_1 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
+interpretation "Exists_elim_ok_1" 'Exists_elim_ok_1 ExPx Pc c =
+    (show c (Exists_elim _ _ _ ExPx Pc)).
+
+notation < "\infrule hbox(\emsp ExPx \emsp\emsp\emsp Pc \emsp) mstyle color #ff0000 (c) (∃\sub\e \emsp ident n \emsp ident HPn)" with precedence 19
+for @{ 'Exists_elim_ok_2 $ExPx (λ${ident n}:$tn.λ${ident HPn}:$Pn.$Pc) $c }.
+interpretation "Exists_elim_ok_2" 'Exists_elim_ok_2 ExPx Pc c =
+    (cast _ _ (show c (Exists_elim _ _ _ ExPx Pc))).
+
+definition ex_concl := λx:CProp.sort → unit → x.
+definition fake_pred := λx:sort.unit.
+
+notation >"∃_'e' term 90 ExPt {ident t} [ident H] term 90 c" non associative with precedence 19
+for @{'Exists_elim (λ_.?) (show $ExPt ?) (λ${ident t}:sort.λ${ident H}.show $c ?)}. 
+interpretation "Exists_elim KO" 'Exists_elim P ExPt c =
+ (cast _ _ (Exists_elim sort P _ 
+   (cast (Exists _ P)  _ ExPt) (cast (ex_concl unit) (ex_concl _) c))).
+interpretation "Exists_elim OK" 'Exists_elim P ExPt c =
+ (Exists_elim sort P _ ExPt c).
+
+(* ∀ introduction *)
+
+notation < "\infrule hbox(\emsp Px \emsp) Pn (mstyle color #ff0000 (∀\sub\i) \emsp ident x)" with precedence 19
+for @{ 'Forall_intro_ko_1 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "Forall_intro_ko_1" 'Forall_intro_ko_1 Px Pn = 
+  (show Pn (cast _ _ (Forall_intro _ _ (cast _ _ Px)))).
+
+notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000(Pn) (mstyle color #ff0000 (∀\sub\i) \emsp ident x)" with precedence 19
+for @{ 'Forall_intro_ko_2 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "Forall_intro_ko_2" 'Forall_intro_ko_2 Px Pn = 
+  (cast _ _ (show Pn (cast _ _ (Forall_intro _ _ (cast _ _ Px))))).
+
+notation < "\infrule hbox(\emsp Px \emsp) Pn (∀\sub\i \emsp ident x)" with precedence 19
+for @{ 'Forall_intro_ok_1 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "Forall_intro_ok_1" 'Forall_intro_ok_1 Px Pn = 
+  (show Pn (Forall_intro _ _ Px)).
+
+notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000 (Pn) (∀\sub\i \emsp ident x)" with precedence 19
+for @{ 'Forall_intro_ok_2 (λ${ident x}:$tx.$Px) $Pn }.
+interpretation "Forall_intro_ok_2" 'Forall_intro_ok_2 Px Pn = 
+  (cast _ _ (show Pn (Forall_intro _ _ Px))).
+
+notation > "∀_'i' {ident y} term 90 Px" non associative with precedence 19
+for @{ 'Forall_intro (λ_.?) (λ${ident y}.show $Px ?) }. 
+interpretation "Forall_intro KO" 'Forall_intro P Px =
+  (cast _ _ (Forall_intro sort P (cast _ _ Px))). 
+interpretation "Forall_intro OK" 'Forall_intro P Px =
+  (Forall_intro sort P Px). 
+
+(* ∀ elimination *)
+notation < "\infrule hbox(\emsp Px \emsp) Pn (mstyle color #ff0000 (∀\sub\e))" with precedence 19
+for @{ 'Forall_elim_ko_1 $Px $Pn }.
+interpretation "Forall_elim_ko_1" 'Forall_elim_ko_1 Px Pn = 
+  (show Pn (cast _ _ (Forall_elim _ _ _ (cast _ _ Px)))).
+
+notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000(Pn) (mstyle color #ff0000 (∀\sub\e))" with precedence 19
+for @{ 'Forall_elim_ko_2 $Px $Pn }.
+interpretation "Forall_elim_ko_2" 'Forall_elim_ko_2 Px Pn = 
+  (cast _ _ (show Pn (cast _ _ (Forall_elim _ _ _ (cast _ _ Px))))).
+
+notation < "\infrule hbox(\emsp Px \emsp) Pn (∀\sub\e)" with precedence 19
+for @{ 'Forall_elim_ok_1 $Px $Pn }.
+interpretation "Forall_elim_ok_1" 'Forall_elim_ok_1 Px Pn = 
+  (show Pn (Forall_elim _ _ _ Px)).
+
+notation < "\infrule hbox(\emsp Px \emsp) mstyle color #ff0000 (Pn) (∀\sub\e)" with precedence 19
+for @{ 'Forall_elim_ok_2 $Px $Pn }.
+interpretation "Forall_elim_ok_2" 'Forall_elim_ok_2 Px Pn = 
+  (cast _ _ (show Pn (Forall_elim _ _ _ Px))).
+
+notation > "∀_'e' {term 90 t} term 90 Pn" non associative with precedence 19
+for @{ 'Forall_elim (λ_.?) $t (show $Pn ?) }. 
+interpretation "Forall_elim KO" 'Forall_elim P t Px =
+  (cast _ unit (Forall_elim sort P t (cast _ _ Px))). 
+interpretation "Forall_elim OK" 'Forall_elim P t Px =
+  (Forall_elim sort P t Px).