una regola, vengono scritti tra `{` e `}`.
* Le formule, quando argomenti di
- una regola, vengono scritti tra `(` e `)`.
+ una regola, vengono scritte tra `(` e `)`.
* I nomi delle ipotesi, quando argomenti di
una regola, vengono scritti tra `[` e `]`.
include "didactic/support/natural_deduction.ma".
-theorem pippo : (∀x.P (f x) ∧ Q x) ⇒ (∃x.Q x) ⇒ ∃y.P y.
-apply rule (prove ((∀x.P (f x) ∧ Q x) ⇒ (∃x.Q x) ⇒ ∃y.P y));
+(* Il nostro linguaggio del primo ordine prevede le seguenti
+ - costanti: c
+ - funzioni: f, g (unarie), h (binaria)
+ - predicati: P, Q (unari), R, S (binari)
+*)
+axiom c : sort.
+axiom f : sort → sort.
+axiom g : sort → sort.
+axiom h : sort → sort → sort.
+axiom P : sort → CProp.
+axiom Q : sort → CProp.
+axiom R : sort →sort → CProp.
+axiom S : sort →sort → CProp.
+
+axiom EM : ∀A:CProp.A ∨ ¬A.
+
+(* intuizionista *)
+lemma ex1: ¬(∃x.P x) ⇒ ∀x.¬ P x.
+apply rule (prove (¬(∃x.P x) ⇒ ∀x.¬ P x));
+apply rule (discharge [x]);
+qed.
+
+(* classico *)
+lemma ex2: ¬(∀x.P x) ⇒ ∃x.¬ P x.
+apply rule (prove (¬(∀x.P x) ⇒ ∃x.¬ P x));
+(*BEGIN*)
+apply rule (⇒_i [h1] (∃x.¬ P x));
+apply rule (RAA [h2] (⊥));
+apply rule (¬_e (¬(∀x.P x)) (∀x.P x));
+ [ apply rule (discharge [h1]);
+ | apply rule (∀_i {y} (P y));
+ apply rule (RAA [h3] (⊥));
+ apply rule (¬_e (¬∃x.¬ P x) (∃x.¬ P x));
+ [ apply rule (discharge [h2]);
+ | apply rule (∃_i {y} (¬P y));
+ apply rule (discharge [h3]);
+ ]
+ ]
+(*END*)
+qed.
+
+(* intuizionista *)
+lemma ex3: ((∃x.P x) ⇒ Q c) ⇒ ∀x.P x ⇒ Q c.
+apply rule (prove (((∃x.P x) ⇒ Q c) ⇒ ∀x.P x ⇒ Q c));
+apply rule (discharge [x]);
+qed.
+
+(* classico *)
+lemma ex4: ((∀x.P x) ⇒ Q c) ⇒ ∃x.P x ⇒ Q c.
+apply rule (prove (((∀x.P x) ⇒ Q c) ⇒ ∃x.P x ⇒ Q c));
(*BEGIN*)
-apply rule (⇒_i [h1] ((∃x.Q x) ⇒ ∃y.P y));
-apply rule (⇒_i [h2] (∃y.P y));
-apply rule (∃_e (∃x.Q x) {t} [h3] (∃y.P y));
- [ apply rule (discharge [h2]);
- | apply rule (let ft ≝ (f t) in ∃_i {ft} (P ft));
- apply rule (∧_e_l (P (f t) ∧ Q t));
- apply rule (∀_e {t} (∀x.P (f x) ∧ Q x));
- apply rule (discharge [h1]);
+apply rule (⇒_i [h1] (∃x.P x ⇒ Q c));
+apply rule (∨_e ((∀x.P x) ∨ ¬(∀x.P x)) [h3] (?) [h3] (?));
+ [ apply rule (lem 0 EM);
+ | apply rule (∃_i {y} (P y ⇒ Q c));
+ apply rule (⇒_i [h4] (Q c));
+ apply rule (⇒_e ((∀x.P x) ⇒ Q c) ((∀x.P x)));
+ [ apply rule (discharge [h1]);
+ | apply rule (discharge [h3]);
+ ]
+ | apply rule (∃_e (∃x.¬P x) {y} [h4] (∃x.P x ⇒ Q c));
+ [ apply rule (lem 1 ex2 (¬(∀x.P x)));
+ apply rule (discharge [h3]);
+ | apply rule (∃_i {y} (P y ⇒ Q c));
+ apply rule (⇒_i [h5] (Q c));
+ apply rule (⊥_e (⊥));
+ apply rule (¬_e (¬P y) (P y));
+ [ apply rule (discharge [h4]);
+ | apply rule (discharge [h5]);
+ ]
+ ]
]
(*END*)
-qed.
\ No newline at end of file
+qed.
+
+
+
+
+
axiom M : CProp.
axiom N : CProp.
axiom O : 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.
+axiom x: sort.
+axiom y: sort.
+axiom z: sort.
+axiom w: sort.
(* Every formula user provided annotates its proof:
`A` becomes `(show A ?)` *)
(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 ?)}.
+for @{'Exists_intro $t (λw.? w) (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 =
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 ex_concl := λx:sort → CProp.∀y:sort.unit → x y.
+definition ex_concl_dummy := ∀y:sort.unit → unit.
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 ?)}.
+for @{'Exists_elim (λx.? x) (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))).
+ (cast (Exists _ P) _ ExPt)
+ (cast ex_concl_dummy (ex_concl _) c))).
interpretation "Exists_elim OK" 'Exists_elim P ExPt c =
(Exists_elim sort P _ ExPt c).