├── .gitignore ├── ClassicalND.v ├── Hilbert.v ├── LICENSE ├── LJTStar.v ├── NDSCEquiv.v ├── NaturalDeduction.v ├── PropLang.v ├── README.md ├── SequentCalculus.v └── Subcontext.v /.gitignore: -------------------------------------------------------------------------------- 1 | Makefile 2 | .*.aux 3 | *.glob 4 | *.v.d 5 | *.vo 6 | html 7 | -------------------------------------------------------------------------------- /ClassicalND.v: -------------------------------------------------------------------------------- 1 | Require Export NaturalDeduction. 2 | Require Import Subcontext. 3 | Require Import NDSCEquiv. 4 | 5 | Section Classical_ND. 6 | 7 | Context {atom : Type}. 8 | 9 | Reserved Notation "Γ ⊢c P" (no associativity, at level 61). 10 | 11 | Inductive classic_ND_proves : list (prop atom) -> prop atom -> Prop := 12 | | classic_ND_exfalso_quodlibet {Γ P} : 13 | Γ ⊢c ⊥ -> 14 | Γ ⊢c P 15 | | classic_ND_True_intro {Γ} : 16 | Γ ⊢c ⊤ 17 | | classic_ND_or_introl {Γ P Q} : 18 | Γ ⊢c P -> 19 | Γ ⊢c P ∨ Q 20 | | classic_ND_or_intror {Γ P Q} : 21 | Γ ⊢c Q -> 22 | Γ ⊢c P ∨ Q 23 | | classic_ND_proof_by_cases {Γ P Q R} : 24 | Γ ⊢c P ∨ Q -> 25 | P :: Γ ⊢c R -> 26 | Q :: Γ ⊢c R -> 27 | Γ ⊢c R 28 | | classic_ND_and_intro {Γ P Q} : 29 | Γ ⊢c P -> 30 | Γ ⊢c Q -> 31 | Γ ⊢c P ∧ Q 32 | | classic_ND_and_elim {Γ P Q R} : 33 | Γ ⊢c P ∧ Q -> 34 | P :: Q :: Γ ⊢c R -> 35 | Γ ⊢c R 36 | | classic_ND_cond_proof {Γ P Q} : 37 | P :: Γ ⊢c Q -> 38 | Γ ⊢c P ⊃ Q 39 | | classic_ND_modus_ponens {Γ P Q} : 40 | Γ ⊢c P ⊃ Q -> 41 | Γ ⊢c P -> 42 | Γ ⊢c Q 43 | | classic_ND_assumption {Γ P} : 44 | In P Γ -> 45 | Γ ⊢c P 46 | | classic_ND_cut {Γ P Q} : 47 | Γ ⊢c P -> 48 | P :: Γ ⊢c Q -> 49 | Γ ⊢c Q 50 | | classic_ND_proof_by_contradiction {Γ P} : 51 | ¬P :: Γ ⊢c ⊥ -> 52 | Γ ⊢c P 53 | where "Γ ⊢c P" := (classic_ND_proves Γ P). 54 | 55 | Proposition ND_impl_classic_ND {Γ P} : 56 | Γ ⊢ P -> Γ ⊢c P. 57 | Proof. 58 | induction 1; 59 | [ econstructor 1 | econstructor 2 | econstructor 3 | 60 | econstructor 4 | econstructor 5 | econstructor 6 | 61 | econstructor 7 | econstructor 8 | econstructor 9 | 62 | econstructor 10 | econstructor 11 ]; eauto. 63 | Qed. 64 | 65 | Proposition classic_ND_NNPP {Γ P} : 66 | Γ ⊢c (¬ ¬ P) ⊃ P. 67 | Proof. 68 | apply @classic_ND_cond_proof. apply @classic_ND_proof_by_contradiction. 69 | apply @classic_ND_modus_ponens with (P := ¬ P); 70 | apply @classic_ND_assumption; unfold not_prop; prove_In. 71 | Qed. 72 | 73 | Proposition classic_ND_excluded_middle {Γ P} : 74 | Γ ⊢c P ∨ ¬ P. 75 | Proof. 76 | apply @classic_ND_proof_by_contradiction. 77 | apply @classic_ND_modus_ponens with (P := P ∨ ¬ P). 78 | + apply @classic_ND_assumption; unfold not_prop; prove_In. 79 | + apply @classic_ND_or_intror. apply @classic_ND_cond_proof. 80 | apply @classic_ND_modus_ponens with (P := P ∨ ¬ P). 81 | - apply @classic_ND_assumption; unfold not_prop; prove_In. 82 | - apply @classic_ND_or_introl. 83 | apply @classic_ND_assumption; prove_In. 84 | Qed. 85 | 86 | Proposition classic_ND_Peirce {Γ P Q} : 87 | Γ ⊢c ((P ⊃ Q) ⊃ P) ⊃ P. 88 | Proof. 89 | apply @classic_ND_cond_proof. 90 | apply @classic_ND_proof_by_contradiction. 91 | apply @classic_ND_modus_ponens with (P := P). 92 | + apply @classic_ND_assumption; unfold not_prop; prove_In. 93 | + apply @classic_ND_modus_ponens with (P := P ⊃ Q). 94 | - apply @classic_ND_assumption; prove_In. 95 | - apply @classic_ND_cond_proof. apply @classic_ND_exfalso_quodlibet. 96 | apply @classic_ND_modus_ponens with (P := P); 97 | apply @classic_ND_assumption; unfold not_prop; prove_In. 98 | Qed. 99 | 100 | Class ND_normal (P : prop atom) : Prop := 101 | | ND_normality : forall Γ, Γ ⊢ (¬ ¬ P) ⊃ P. 102 | 103 | Lemma ND_double_neg_elim_for_normal {Γ P Q} `{ND_normal Q} : 104 | Γ ⊢ (¬ ¬ P) -> P :: Γ ⊢ Q -> Γ ⊢ Q. 105 | Proof. 106 | intros. apply @ND_modus_ponens with (P := ¬ ¬ Q). 107 | + apply H. 108 | + apply @ND_cond_proof. apply @ND_cut with (P := ¬ ¬ P). 109 | - eapply @ND_context_extension with (3 := H0); 110 | (prove_subcontext || reflexivity). 111 | - apply @ND_modus_ponens with (P := ¬ P). 112 | * apply @ND_assumption; unfold not_prop; prove_In. 113 | * apply @ND_cond_proof. apply @ND_modus_ponens with (P := Q). 114 | { apply @ND_assumption; unfold not_prop; prove_In. } 115 | { eapply @ND_context_extension with (3 := H1); 116 | (prove_subcontext || reflexivity). } 117 | Qed. 118 | 119 | Instance ND_normal_bot : ND_normal ⊥. 120 | Proof. 121 | intro Γ. unfold not_prop. apply @ND_cond_proof. 122 | apply @ND_modus_ponens with (P := ⊥ ⊃ ⊥). 123 | + apply @ND_assumption; prove_In. 124 | + apply @ND_cond_proof. apply @ND_assumption; prove_In. 125 | Qed. 126 | 127 | Instance ND_normal_impl {P Q} : ND_normal Q -> ND_normal (P ⊃ Q). 128 | Proof. 129 | intros ? Γ. do 2 apply @ND_cond_proof. 130 | apply @ND_double_neg_elim_for_normal with (P := P ⊃ Q); auto. 131 | + apply @ND_assumption; prove_In. 132 | + apply @ND_modus_ponens with (P := P); 133 | apply @ND_assumption; prove_In. 134 | Qed. 135 | 136 | Theorem classic_ND_equiv {Γ P} : 137 | Γ ⊢c P <-> Γ ⊢ ¬ (¬ P). 138 | Proof. 139 | split; intros. 140 | + induction H. 141 | - apply @ND_exfalso_quodlibet. apply @ND_modus_ponens with (P := ¬ ⊥). 142 | * assumption. 143 | * apply @ND_cond_proof. apply @ND_assumption; prove_In. 144 | - apply @ND_cond_proof. apply @ND_modus_ponens with (P := ⊤). 145 | * apply @ND_assumption; unfold not_prop; prove_In. 146 | * apply @ND_True_intro. 147 | - apply (ND_double_neg_elim_for_normal IHclassic_ND_proves). 148 | apply @ND_cond_proof. apply @ND_modus_ponens with (P := P ∨ Q). 149 | * apply @ND_assumption; unfold not_prop; prove_In. 150 | * apply @ND_or_introl. apply @ND_assumption; prove_In. 151 | - apply (ND_double_neg_elim_for_normal IHclassic_ND_proves). 152 | apply @ND_cond_proof. apply @ND_modus_ponens with (P := P ∨ Q). 153 | * apply @ND_assumption; unfold not_prop; prove_In. 154 | * apply @ND_or_intror. apply @ND_assumption; prove_In. 155 | - apply (ND_double_neg_elim_for_normal IHclassic_ND_proves1). 156 | apply @ND_proof_by_cases with (P := P) (Q := Q). 157 | * apply @ND_assumption; prove_In. 158 | * apply @ND_context_extension with (3 := IHclassic_ND_proves2); 159 | (prove_subcontext || reflexivity). 160 | * apply @ND_context_extension with (3 := IHclassic_ND_proves3); 161 | (prove_subcontext || reflexivity). 162 | - apply (ND_double_neg_elim_for_normal IHclassic_ND_proves1). 163 | apply (ND_double_neg_elim_for_normal (P := Q)). 164 | * apply @ND_context_extension with (3 := IHclassic_ND_proves2); 165 | (prove_subcontext || reflexivity). 166 | * apply @ND_cond_proof. apply @ND_modus_ponens with (P := P ∧ Q). 167 | { apply @ND_assumption; unfold not_prop; prove_In. } 168 | { apply @ND_and_intro; apply @ND_assumption; prove_In. } 169 | - apply (ND_double_neg_elim_for_normal IHclassic_ND_proves1). 170 | apply @ND_and_elim with (P := P) (Q := Q). 171 | * apply @ND_assumption; prove_In. 172 | * apply @ND_context_extension with (3 := IHclassic_ND_proves2); 173 | (prove_subcontext || reflexivity). 174 | - apply @ND_cond_proof. apply @ND_modus_ponens with (P := P ⊃ Q). 175 | * apply @ND_assumption; unfold not_prop; prove_In. 176 | * apply @ND_cond_proof. apply @ND_cut with (P := ¬ ¬ Q). 177 | { apply @ND_context_extension with (3 := IHclassic_ND_proves); 178 | (prove_subcontext || reflexivity). } 179 | { apply @ND_exfalso_quodlibet. apply @ND_modus_ponens with (P := ¬ Q). 180 | + apply @ND_assumption; unfold not_prop; prove_In. 181 | + apply @ND_cond_proof. apply @ND_modus_ponens with (P := P ⊃ Q). 182 | - apply @ND_assumption; unfold not_prop; prove_In. 183 | - apply @ND_cond_proof. apply @ND_assumption; prove_In. } 184 | - apply (ND_double_neg_elim_for_normal IHclassic_ND_proves1). 185 | apply (ND_double_neg_elim_for_normal (P := P)). 186 | * apply @ND_context_extension with (3 := IHclassic_ND_proves2); 187 | (prove_subcontext || reflexivity). 188 | * apply @ND_cond_proof. apply @ND_modus_ponens with (P := Q). 189 | { apply @ND_assumption; unfold not_prop; prove_In. } 190 | { apply @ND_modus_ponens with (P := P); 191 | apply @ND_assumption; prove_In. } 192 | - apply @ND_cond_proof. apply @ND_modus_ponens with (P := P). 193 | * apply @ND_assumption; unfold not_prop; prove_In. 194 | * apply @ND_assumption; right; assumption. 195 | - apply (ND_double_neg_elim_for_normal IHclassic_ND_proves1). 196 | assumption. 197 | - apply @ND_cond_proof. apply @ND_modus_ponens with (P := ¬ ⊥). 198 | * assumption. 199 | * apply @ND_cond_proof. apply @ND_assumption; prove_In. 200 | + apply @classic_ND_modus_ponens with (P := ¬ (¬ P)); 201 | auto using @ND_impl_classic_ND, @classic_ND_NNPP. 202 | Qed. 203 | 204 | Corollary classic_ND_consistent : ~ (nil ⊢c ⊥). 205 | Proof. 206 | rewrite classic_ND_equiv. intro. eapply @ND_consistent. 207 | eapply @ND_modus_ponens with (P := ¬ ⊥). 208 | + eassumption. 209 | + apply @ND_cond_proof. apply @ND_assumption; prove_In. 210 | Qed. 211 | 212 | End Classical_ND. 213 | -------------------------------------------------------------------------------- /Hilbert.v: -------------------------------------------------------------------------------- 1 | Require Export PropLang. 2 | Require Export List. 3 | Require Export Subcontext. 4 | Require Export RelationClasses. 5 | Require Export Morphisms. 6 | Require Export NaturalDeduction. 7 | 8 | Section Hilbert. 9 | 10 | Context {atom : Type}. 11 | 12 | Inductive hilbert_axiom : prop atom -> Prop := 13 | | hilbert_axiom_I {P} : hilbert_axiom (P ⊃ P) 14 | | hilbert_axiom_K {P Q} : hilbert_axiom (P ⊃ Q ⊃ P) 15 | | hilbert_axiom_S {P Q R} : hilbert_axiom ((P ⊃ Q ⊃ R) ⊃ (P ⊃ Q) ⊃ P ⊃ R) 16 | | hilbert_axiom_bot_elim {P} : hilbert_axiom (⊥ ⊃ P) 17 | | hilbert_axiom_top_intro : hilbert_axiom ⊤ 18 | | hilbert_axiom_and_intro {P Q} : hilbert_axiom (P ⊃ Q ⊃ P ∧ Q) 19 | | hilbert_axiom_and_elim {P Q R} : hilbert_axiom ((P ⊃ Q ⊃ R) ⊃ P ∧ Q ⊃ R) 20 | | hilbert_axiom_or_introl {P Q} : hilbert_axiom (P ⊃ P ∨ Q) 21 | | hilbert_axiom_or_intror {P Q} : hilbert_axiom (Q ⊃ P ∨ Q) 22 | | hilbert_axiom_or_elim {P Q R} : hilbert_axiom ((P ⊃ R) ⊃ (Q ⊃ R) ⊃ P ∨ Q ⊃ R). 23 | 24 | Reserved Notation "Γ ≤h Γ'" (no associativity, at level 61). 25 | 26 | (* Γ ≤h Γ' means that there is a valid sequence of steps starting 27 | from Γ with the additional steps along with Γ giving Γ'. 28 | Note that for convenience we add each step at the front of the 29 | list instead of the back. *) 30 | Inductive hilbert_derivation : list (prop atom) -> list (prop atom) -> Prop := 31 | | hilbert_empty {Γ} : Γ ≤h Γ 32 | | hilbert_axiom_derivation {Γ Γ' P} : Γ ≤h Γ' -> hilbert_axiom P -> 33 | Γ ≤h P :: Γ' 34 | | hilbert_modus_ponens {Γ Γ' P Q} : Γ ≤h Γ' -> In P Γ' -> In (P ⊃ Q) Γ' -> 35 | Γ ≤h Q :: Γ' 36 | where "Γ ≤h Γ'" := (hilbert_derivation Γ Γ'). 37 | 38 | Global Instance hilbert_derivation_preord : PreOrder hilbert_derivation. 39 | Proof. 40 | constructor. 41 | + intro Γ; constructor. 42 | + intros Γ Γ' Γ'' H H0. induction H0. 43 | - assumption. 44 | - constructor 2; auto. 45 | - constructor 3 with (P := P); eauto. 46 | Qed. 47 | 48 | Lemma hilbert_derivation_tail : forall Γ Γ', Γ ≤h Γ' -> 49 | exists Γ'', Γ' = Γ'' ++ Γ. 50 | Proof. 51 | induction 1. 52 | + exists nil; reflexivity. 53 | + destruct IHhilbert_derivation as [Γ'']. exists (P :: Γ''). 54 | rewrite H1. reflexivity. 55 | + destruct IHhilbert_derivation as [Γ'']. exists (Q :: Γ''). 56 | rewrite H2. reflexivity. 57 | Qed. 58 | 59 | Lemma hilbert_derivation_context_extension : 60 | forall Γ Γ' Γ'', Γ ⊆ Γ' -> Γ ≤h Γ'' ++ Γ -> 61 | Γ' ≤h Γ'' ++ Γ'. 62 | Proof. 63 | intros. remember (Γ'' ++ Γ) as Γ₀. revert Γ'' HeqΓ₀. induction H0. 64 | + intros. change (nil ++ Γ = Γ'' ++ Γ) in HeqΓ₀. 65 | apply app_inv_tail in HeqΓ₀. subst; simpl. constructor. 66 | + intros. destruct (hilbert_derivation_tail _ _ H0) as [Γh]. subst. 67 | pose proof (IHhilbert_derivation H Γh eq_refl). 68 | change ((P :: Γh) ++ Γ = Γ'' ++ Γ) in HeqΓ₀. 69 | apply app_inv_tail in HeqΓ₀. subst. simpl. constructor 2; auto. 70 | + intros. destruct (hilbert_derivation_tail _ _ H0) as [Γh]. subst. 71 | pose proof (IHhilbert_derivation H Γh eq_refl). 72 | change ((Q :: Γh) ++ Γ = Γ'' ++ Γ) in HeqΓ₀. 73 | apply app_inv_tail in HeqΓ₀. subst. simpl. 74 | constructor 3 with (P := P); auto. 75 | - destruct (in_app_or _ _ _ H1). 76 | * apply in_or_app; left; assumption. 77 | * apply in_or_app; right; auto. 78 | - destruct (in_app_or _ _ _ H2). 79 | * apply in_or_app; left; assumption. 80 | * apply in_or_app; right; auto. 81 | Qed. 82 | 83 | Inductive hilbert_proves (Γ : list (prop atom)) (P : prop atom) : Prop := 84 | | hilbert_derivation_proves : forall Γ', Γ ≤h (P :: Γ') -> hilbert_proves Γ P. 85 | Notation "Γ ⊢h P" := (hilbert_proves Γ P) (no associativity, at level 61). 86 | 87 | Proposition hilbert_assumption {Γ P} : In P Γ -> Γ ⊢h P. 88 | Proof. 89 | intros. apply hilbert_derivation_proves with (Γ' := P ⊃ P :: Γ). 90 | apply @hilbert_modus_ponens with (P := P). 91 | + apply hilbert_axiom_derivation. 92 | - constructor. 93 | - constructor. 94 | + right; assumption. 95 | + left; reflexivity. 96 | Qed. 97 | 98 | Global Instance hilbert_context_extension : 99 | Proper (subcontext ++> eq ==> Basics.impl) hilbert_proves. 100 | Proof. 101 | intros Γ Γ' ? P Q [] ?. destruct H0. destruct (hilbert_derivation_tail _ _ H0). 102 | destruct x. 103 | + simpl in H1. subst. apply hilbert_assumption. apply H. prove_In. 104 | + injection H1; intros; subst. change (Γ ≤h (p :: x) ++ Γ) in H0. 105 | pose proof (hilbert_derivation_context_extension _ _ _ H H0). 106 | eauto using hilbert_derivation_proves. 107 | Qed. 108 | 109 | Proposition hilbert_cut {Γ P Q} : 110 | Γ ⊢h P -> P :: Γ ⊢h Q -> Γ ⊢h Q. 111 | Proof. 112 | destruct 1. intros. 113 | simple refine (let H1 := hilbert_context_extension (P :: Γ) (P :: Γ') _ _ _ eq_refl H0 in _). 114 | + rewrite subcontext_cons; split. 115 | - prove_In. 116 | - destruct (hilbert_derivation_tail _ _ H). red; intros. 117 | rewrite H1. apply in_or_app; right; assumption. 118 | + destruct H1. rewrite <- H in H1. eauto using hilbert_derivation_proves. 119 | Qed. 120 | 121 | Inductive hilbert_impl_lift (Γ Γ' : list (prop atom)) (P : prop atom) : Prop := 122 | | hilbert_impl_lift_intro : forall Γ'', Γ ≤h Γ'' -> 123 | (forall Q, In Q Γ' -> In (P ⊃ Q) Γ'') -> hilbert_impl_lift Γ Γ' P. 124 | 125 | Lemma hilbert_cond_derivation : forall Γ Γ' P, P :: Γ ≤h Γ' -> 126 | hilbert_impl_lift Γ Γ' P. 127 | Proof. 128 | intros. remember (P :: Γ) as PΓ. revert P Γ HeqPΓ. 129 | induction H; intros; subst. 130 | + assert (forall Γ, Γ ⊆ Γ0 -> hilbert_impl_lift Γ0 Γ P). 131 | - induction Γ. 132 | * exists Γ0. 133 | { constructor. } 134 | { destruct 1. } 135 | * rewrite subcontext_cons. destruct 1. destruct (IHΓ H0). 136 | assert (Γ0 ≤h a ⊃ P ⊃ a :: Γ'') by 137 | eauto using hilbert_derivation, hilbert_axiom_K. 138 | assert (Γ0 ≤h P ⊃ a :: a ⊃ P ⊃ a :: Γ''). 139 | { apply @hilbert_modus_ponens with (P := a); auto; try prove_In. 140 | right. destruct (hilbert_derivation_tail _ _ H1). 141 | rewrite H4. apply in_or_app; right; assumption. } 142 | { apply hilbert_impl_lift_intro with (1 := H4). 143 | destruct 1; subst; (prove_In || (do 2 right; auto)). } 144 | - destruct (H Γ0); try reflexivity. 145 | assert (Γ0 ≤h P ⊃ P :: Γ'') by 146 | eauto using hilbert_derivation, hilbert_axiom_I. 147 | apply hilbert_impl_lift_intro with (1 := H2). 148 | destruct 1; subst; (prove_In || (right; auto)). 149 | + destruct (IHhilbert_derivation _ _ eq_refl). 150 | assert (Γ0 ≤h P :: Γ'') by eauto using hilbert_derivation. 151 | assert (Γ0 ≤h P ⊃ P0 ⊃ P :: P :: Γ'') by eauto using 152 | hilbert_derivation, hilbert_axiom_K. 153 | assert (Γ0 ≤h P0 ⊃ P :: P ⊃ P0 ⊃ P :: P :: Γ'') by 154 | (apply @hilbert_modus_ponens with (P := P); auto; prove_In). 155 | apply hilbert_impl_lift_intro with (1 := H5). 156 | destruct 1; subst; (prove_In || (do 3 right; auto)). 157 | + destruct (IHhilbert_derivation _ _ eq_refl). 158 | pose proof (H3 _ H0); pose proof (H3 _ H1). 159 | assert (Γ0 ≤h (P0 ⊃ P ⊃ Q) ⊃ (P0 ⊃ P) ⊃ (P0 ⊃ Q) :: Γ'') by 160 | eauto using hilbert_derivation, hilbert_axiom_S. 161 | assert (Γ0 ≤h (P0 ⊃ P) ⊃ P0 ⊃ Q :: 162 | (P0 ⊃ P ⊃ Q) ⊃ (P0 ⊃ P) ⊃ (P0 ⊃ Q) :: Γ'') by 163 | (apply @hilbert_modus_ponens with (P := P0 ⊃ P ⊃ Q); auto; 164 | (prove_In || (right; auto))). 165 | assert (Γ0 ≤h P0 ⊃ Q :: (P0 ⊃ P) ⊃ P0 ⊃ Q :: 166 | (P0 ⊃ P ⊃ Q) ⊃ (P0 ⊃ P) ⊃ (P0 ⊃ Q) :: Γ'') by 167 | (apply @hilbert_modus_ponens with (P := P0 ⊃ P); auto; 168 | (prove_In || (right; right; auto))). 169 | apply hilbert_impl_lift_intro with (1 := H8). 170 | destruct 1; subst; (prove_In || (do 3 right; auto)). 171 | Qed. 172 | 173 | Theorem hilbert_cond_proof {P Q Γ} : 174 | P :: Γ ⊢h Q -> Γ ⊢h P ⊃ Q. 175 | Proof. 176 | intros. destruct H. destruct (hilbert_cond_derivation _ _ _ H). 177 | assert (In (P ⊃ Q) Γ'') by (apply H1; prove_In). 178 | destruct (hilbert_assumption H2). rewrite <- H0 in H3. 179 | eauto using hilbert_derivation_proves. 180 | Qed. 181 | 182 | 183 | Proposition hilbert_axiom_soundness {Γ P} : 184 | hilbert_axiom P -> Γ ⊢ P. 185 | Proof. 186 | destruct 1. 187 | + apply ND_cond_proof. apply ND_assumption. prove_In. 188 | + apply ND_cond_proof. apply ND_cond_proof. apply ND_assumption. prove_In. 189 | + do 3 apply ND_cond_proof. apply @ND_modus_ponens with (P := Q). 190 | - apply @ND_modus_ponens with (P := P); apply ND_assumption; prove_In. 191 | - apply @ND_modus_ponens with (P := P); apply ND_assumption; prove_In. 192 | + apply ND_cond_proof. apply ND_exfalso_quodlibet. 193 | apply ND_assumption; prove_In. 194 | + apply ND_True_intro. 195 | + do 2 apply ND_cond_proof. apply ND_and_intro; apply ND_assumption; prove_In. 196 | + do 2 apply ND_cond_proof. apply @ND_and_elim with (P := P) (Q := Q). 197 | - apply ND_assumption; prove_In. 198 | - apply @ND_modus_ponens with (P := Q). 199 | * apply @ND_modus_ponens with (P := P); apply ND_assumption; prove_In. 200 | * apply ND_assumption; prove_In. 201 | + apply ND_cond_proof. apply ND_or_introl; apply ND_assumption; prove_In. 202 | + apply ND_cond_proof. apply ND_or_intror; apply ND_assumption; prove_In. 203 | + do 3 apply ND_cond_proof. apply @ND_proof_by_cases with (P := P) (Q := Q). 204 | - apply ND_assumption; prove_In. 205 | - apply @ND_modus_ponens with (P := P); apply ND_assumption; prove_In. 206 | - apply @ND_modus_ponens with (P := Q); apply ND_assumption; prove_In. 207 | Qed. 208 | 209 | Proposition hilbert_derivation_soundness {Γ Γ'} : 210 | Γ ≤h Γ' -> forall P, In P Γ' -> Γ ⊢ P. 211 | Proof. 212 | induction 1. 213 | + apply @ND_assumption. 214 | + destruct 1; subst; auto using hilbert_axiom_soundness. 215 | + destruct 1; subst; auto. 216 | apply @ND_modus_ponens with (P := P); auto. 217 | Qed. 218 | 219 | Theorem hilbert_soundness {Γ P} : Γ ⊢h P -> Γ ⊢ P. 220 | Proof. 221 | destruct 1. refine (hilbert_derivation_soundness H _ _). prove_In. 222 | Qed. 223 | 224 | Proposition hilbert_proves_axiom {Γ P} : hilbert_axiom P -> Γ ⊢h P. 225 | Proof. 226 | intros. exists Γ. eauto using hilbert_derivation. 227 | Qed. 228 | 229 | Theorem hilbert_completeness {Γ P} : Γ ⊢ P -> Γ ⊢h P. 230 | Proof. 231 | induction 1. 232 | + destruct IHND_proves. assert (Γ ≤h ⊥ ⊃ P :: ⊥ :: Γ') by 233 | eauto using hilbert_derivation, hilbert_axiom_bot_elim. 234 | assert (Γ ≤h P :: ⊥ ⊃ P :: ⊥ :: Γ') by 235 | (apply @hilbert_modus_ponens with (P := ⊥); auto; prove_In). 236 | eauto using hilbert_derivation_proves. 237 | + apply hilbert_proves_axiom. constructor. 238 | + destruct IHND_proves. assert (Γ ≤h P ⊃ P ∨ Q :: P :: Γ') by 239 | eauto using hilbert_derivation, hilbert_axiom_or_introl. 240 | assert (Γ ≤h P ∨ Q :: P ⊃ P ∨ Q :: P :: Γ') by 241 | (apply @hilbert_modus_ponens with (P := P); auto; prove_In). 242 | eauto using hilbert_derivation_proves. 243 | + destruct IHND_proves. assert (Γ ≤h Q ⊃ P ∨ Q :: Q :: Γ') by 244 | eauto using hilbert_derivation, hilbert_axiom_or_intror. 245 | assert (Γ ≤h P ∨ Q :: Q ⊃ P ∨ Q :: Q :: Γ') by 246 | (apply @hilbert_modus_ponens with (P := Q); auto; prove_In). 247 | eauto using hilbert_derivation_proves. 248 | + apply hilbert_cond_proof in IHND_proves2; 249 | apply hilbert_cond_proof in IHND_proves3. 250 | apply (hilbert_cut IHND_proves1). 251 | apply @hilbert_cut with (P := P ⊃ R). 252 | - refine (hilbert_context_extension _ _ _ _ _ eq_refl IHND_proves2). 253 | prove_subcontext. 254 | - apply @hilbert_cut with (P := Q ⊃ R). 255 | * refine (hilbert_context_extension _ _ _ _ _ eq_refl IHND_proves3). 256 | prove_subcontext. 257 | * set (Γ' := Q ⊃ R :: P ⊃ R :: P ∨ Q :: Γ). 258 | assert (Γ' ≤h (P ⊃ R) ⊃ (Q ⊃ R) ⊃ P ∨ Q ⊃ R :: Γ') by 259 | eauto using hilbert_derivation, hilbert_axiom_or_elim. 260 | assert (Γ' ≤h (Q ⊃ R) ⊃ P ∨ Q ⊃ R :: (P ⊃ R) ⊃ (Q ⊃ R) ⊃ P ∨ Q ⊃ R :: Γ') by 261 | (apply @hilbert_modus_ponens with (P := P ⊃ R); auto; 262 | unfold Γ'; prove_In). 263 | assert (Γ' ≤h P ∨ Q ⊃ R :: (Q ⊃ R) ⊃ P ∨ Q ⊃ R :: 264 | (P ⊃ R) ⊃ (Q ⊃ R) ⊃ P ∨ Q ⊃ R :: Γ') by 265 | (apply @hilbert_modus_ponens with (P := Q ⊃ R); auto; 266 | unfold Γ'; prove_In). 267 | assert (Γ' ≤h R :: P ∨ Q ⊃ R :: (Q ⊃ R) ⊃ P ∨ Q ⊃ R :: 268 | (P ⊃ R) ⊃ (Q ⊃ R) ⊃ P ∨ Q ⊃ R :: Γ') by 269 | (apply @hilbert_modus_ponens with (P := P ∨ Q); auto; 270 | unfold Γ'; prove_In). 271 | eauto using hilbert_derivation_proves. 272 | + apply (hilbert_cut IHND_proves1). 273 | apply @hilbert_cut with (P := Q). 274 | - refine (hilbert_context_extension _ _ _ _ _ eq_refl IHND_proves2). 275 | prove_subcontext. 276 | - set (Γ' := Q :: P :: Γ). 277 | assert (Γ' ≤h P ⊃ Q ⊃ P ∧ Q :: Γ') by 278 | eauto using hilbert_derivation, hilbert_axiom_and_intro. 279 | assert (Γ' ≤h Q ⊃ P ∧ Q :: P ⊃ Q ⊃ P ∧ Q :: Γ') by 280 | (apply @hilbert_modus_ponens with (P := P); auto; 281 | unfold Γ'; prove_In). 282 | assert (Γ' ≤h P ∧ Q :: Q ⊃ P ∧ Q :: P ⊃ Q ⊃ P ∧ Q :: Γ') by 283 | (apply @hilbert_modus_ponens with (P := Q); auto; 284 | unfold Γ'; prove_In). 285 | eauto using hilbert_derivation_proves. 286 | + apply (hilbert_cut IHND_proves1). 287 | apply @hilbert_cut with (P := P ⊃ Q ⊃ R). 288 | - do 2 apply hilbert_cond_proof. 289 | refine (hilbert_context_extension _ _ _ _ _ eq_refl IHND_proves2). 290 | prove_subcontext. 291 | - set (Γ' := P ⊃ Q ⊃ R :: P ∧ Q :: Γ). 292 | assert (Γ' ≤h (P ⊃ Q ⊃ R) ⊃ P ∧ Q ⊃ R :: Γ') by 293 | eauto using hilbert_derivation, hilbert_axiom_and_elim. 294 | assert (Γ' ≤h P ∧ Q ⊃ R :: (P ⊃ Q ⊃ R) ⊃ P ∧ Q ⊃ R :: Γ') by 295 | (apply @hilbert_modus_ponens with (P := P ⊃ Q ⊃ R); auto; 296 | unfold Γ'; prove_In). 297 | assert (Γ' ≤h R :: P ∧ Q ⊃ R :: (P ⊃ Q ⊃ R) ⊃ P ∧ Q ⊃ R :: Γ') by 298 | (apply @hilbert_modus_ponens with (P := P ∧ Q); auto; 299 | unfold Γ'; prove_In). 300 | eauto using hilbert_derivation_proves. 301 | + auto using hilbert_cond_proof. 302 | + apply (hilbert_cut IHND_proves1). 303 | apply @hilbert_cut with (P := P). 304 | - refine (hilbert_context_extension _ _ _ _ _ eq_refl IHND_proves2). 305 | prove_subcontext. 306 | - set (Γ' := P :: P ⊃ Q :: Γ). exists Γ'. 307 | apply @hilbert_modus_ponens with (P := P); (apply hilbert_empty || 308 | (unfold Γ'; prove_In)). 309 | + auto using hilbert_assumption. 310 | + eauto using hilbert_cut. 311 | Qed. 312 | 313 | Theorem ND_hilbert_equiv : forall Γ P, Γ ⊢ P <-> Γ ⊢h P. 314 | Proof. 315 | split; [ apply hilbert_completeness | apply hilbert_soundness ]. 316 | Qed. 317 | 318 | End Hilbert. 319 | 320 | Notation "Γ ≤h Γ'" := (hilbert_derivation Γ Γ') (no associativity, at level 61). 321 | Notation "Γ ⊢h P" := (hilbert_proves Γ P) (no associativity, at level 61). 322 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | GNU LESSER GENERAL PUBLIC LICENSE 2 | Version 3, 29 June 2007 3 | 4 | Copyright (C) 2007 Free Software Foundation, Inc. 5 | Everyone is permitted to copy and distribute verbatim copies 6 | of this license document, but changing it is not allowed. 7 | 8 | 9 | This version of the GNU Lesser General Public License incorporates 10 | the terms and conditions of version 3 of the GNU General Public 11 | License, supplemented by the additional permissions listed below. 12 | 13 | 0. Additional Definitions. 14 | 15 | As used herein, "this License" refers to version 3 of the GNU Lesser 16 | General Public License, and the "GNU GPL" refers to version 3 of the GNU 17 | General Public License. 18 | 19 | "The Library" refers to a covered work governed by this License, 20 | other than an Application or a Combined Work as defined below. 21 | 22 | An "Application" is any work that makes use of an interface provided 23 | by the Library, but which is not otherwise based on the Library. 24 | Defining a subclass of a class defined by the Library is deemed a mode 25 | of using an interface provided by the Library. 26 | 27 | A "Combined Work" is a work produced by combining or linking an 28 | Application with the Library. The particular version of the Library 29 | with which the Combined Work was made is also called the "Linked 30 | Version". 31 | 32 | The "Minimal Corresponding Source" for a Combined Work means the 33 | Corresponding Source for the Combined Work, excluding any source code 34 | for portions of the Combined Work that, considered in isolation, are 35 | based on the Application, and not on the Linked Version. 36 | 37 | The "Corresponding Application Code" for a Combined Work means the 38 | object code and/or source code for the Application, including any data 39 | and utility programs needed for reproducing the Combined Work from the 40 | Application, but excluding the System Libraries of the Combined Work. 41 | 42 | 1. Exception to Section 3 of the GNU GPL. 43 | 44 | You may convey a covered work under sections 3 and 4 of this License 45 | without being bound by section 3 of the GNU GPL. 46 | 47 | 2. Conveying Modified Versions. 48 | 49 | If you modify a copy of the Library, and, in your modifications, a 50 | facility refers to a function or data to be supplied by an Application 51 | that uses the facility (other than as an argument passed when the 52 | facility is invoked), then you may convey a copy of the modified 53 | version: 54 | 55 | a) under this License, provided that you make a good faith effort to 56 | ensure that, in the event an Application does not supply the 57 | function or data, the facility still operates, and performs 58 | whatever part of its purpose remains meaningful, or 59 | 60 | b) under the GNU GPL, with none of the additional permissions of 61 | this License applicable to that copy. 62 | 63 | 3. Object Code Incorporating Material from Library Header Files. 64 | 65 | The object code form of an Application may incorporate material from 66 | a header file that is part of the Library. You may convey such object 67 | code under terms of your choice, provided that, if the incorporated 68 | material is not limited to numerical parameters, data structure 69 | layouts and accessors, or small macros, inline functions and templates 70 | (ten or fewer lines in length), you do both of the following: 71 | 72 | a) Give prominent notice with each copy of the object code that the 73 | Library is used in it and that the Library and its use are 74 | covered by this License. 75 | 76 | b) Accompany the object code with a copy of the GNU GPL and this license 77 | document. 78 | 79 | 4. Combined Works. 80 | 81 | You may convey a Combined Work under terms of your choice that, 82 | taken together, effectively do not restrict modification of the 83 | portions of the Library contained in the Combined Work and reverse 84 | engineering for debugging such modifications, if you also do each of 85 | the following: 86 | 87 | a) Give prominent notice with each copy of the Combined Work that 88 | the Library is used in it and that the Library and its use are 89 | covered by this License. 90 | 91 | b) Accompany the Combined Work with a copy of the GNU GPL and this license 92 | document. 93 | 94 | c) For a Combined Work that displays copyright notices during 95 | execution, include the copyright notice for the Library among 96 | these notices, as well as a reference directing the user to the 97 | copies of the GNU GPL and this license document. 98 | 99 | d) Do one of the following: 100 | 101 | 0) Convey the Minimal Corresponding Source under the terms of this 102 | License, and the Corresponding Application Code in a form 103 | suitable for, and under terms that permit, the user to 104 | recombine or relink the Application with a modified version of 105 | the Linked Version to produce a modified Combined Work, in the 106 | manner specified by section 6 of the GNU GPL for conveying 107 | Corresponding Source. 108 | 109 | 1) Use a suitable shared library mechanism for linking with the 110 | Library. A suitable mechanism is one that (a) uses at run time 111 | a copy of the Library already present on the user's computer 112 | system, and (b) will operate properly with a modified version 113 | of the Library that is interface-compatible with the Linked 114 | Version. 115 | 116 | e) Provide Installation Information, but only if you would otherwise 117 | be required to provide such information under section 6 of the 118 | GNU GPL, and only to the extent that such information is 119 | necessary to install and execute a modified version of the 120 | Combined Work produced by recombining or relinking the 121 | Application with a modified version of the Linked Version. (If 122 | you use option 4d0, the Installation Information must accompany 123 | the Minimal Corresponding Source and Corresponding Application 124 | Code. If you use option 4d1, you must provide the Installation 125 | Information in the manner specified by section 6 of the GNU GPL 126 | for conveying Corresponding Source.) 127 | 128 | 5. Combined Libraries. 129 | 130 | You may place library facilities that are a work based on the 131 | Library side by side in a single library together with other library 132 | facilities that are not Applications and are not covered by this 133 | License, and convey such a combined library under terms of your 134 | choice, if you do both of the following: 135 | 136 | a) Accompany the combined library with a copy of the same work based 137 | on the Library, uncombined with any other library facilities, 138 | conveyed under the terms of this License. 139 | 140 | b) Give prominent notice with the combined library that part of it 141 | is a work based on the Library, and explaining where to find the 142 | accompanying uncombined form of the same work. 143 | 144 | 6. Revised Versions of the GNU Lesser General Public License. 145 | 146 | The Free Software Foundation may publish revised and/or new versions 147 | of the GNU Lesser General Public License from time to time. Such new 148 | versions will be similar in spirit to the present version, but may 149 | differ in detail to address new problems or concerns. 150 | 151 | Each version is given a distinguishing version number. If the 152 | Library as you received it specifies that a certain numbered version 153 | of the GNU Lesser General Public License "or any later version" 154 | applies to it, you have the option of following the terms and 155 | conditions either of that published version or of any later version 156 | published by the Free Software Foundation. If the Library as you 157 | received it does not specify a version number of the GNU Lesser 158 | General Public License, you may choose any version of the GNU Lesser 159 | General Public License ever published by the Free Software Foundation. 160 | 161 | If the Library as you received it specifies that a proxy can decide 162 | whether future versions of the GNU Lesser General Public License shall 163 | apply, that proxy's public statement of acceptance of any version is 164 | permanent authorization for you to choose that version for the 165 | Library. 166 | -------------------------------------------------------------------------------- /LJTStar.v: -------------------------------------------------------------------------------- 1 | Require Export NaturalDeduction. 2 | Require Import SequentCalculus. 3 | Require Import NDSCEquiv. 4 | 5 | Section LJTstar. 6 | 7 | Context {atom : Type}. 8 | 9 | Reserved Notation "Γ ⇒* P" (no associativity, at level 61). 10 | 11 | Inductive remove (x : prop atom) : list (prop atom) -> list (prop atom) -> Prop := 12 | | remove_here {l} : remove x (x :: l) l 13 | | remove_later {h l l'} : remove x l l' -> remove x (h :: l) (h :: l'). 14 | 15 | Lemma In_ex_remove : forall x l, In x l <-> exists l', remove x l l'. 16 | Proof. 17 | split; intros. 18 | + induction l; destruct H; subst. 19 | - exists l; constructor. 20 | - destruct (IHl H) as [l0]. exists (a :: l0); constructor; auto. 21 | + destruct H as [l']. induction H. 22 | - left; reflexivity. 23 | - right; assumption. 24 | Qed. 25 | 26 | Corollary remove_In : forall x l l', remove x l l' -> In x l. 27 | Proof. 28 | intros. rewrite In_ex_remove. eauto. 29 | Qed. 30 | 31 | Lemma In_remove : forall y x l l', remove x l l' -> 32 | (In y l <-> y = x \/ In y l'). 33 | Proof. 34 | induction 1; intuition. 35 | + destruct H; intuition. 36 | + left. symmetry; assumption. 37 | + destruct H1; subst. 38 | - right; left; reflexivity. 39 | - destruct (H0 H1); intuition. 40 | + destruct H4; subst. 41 | - left; reflexivity. 42 | - right; auto. 43 | Qed. 44 | 45 | Lemma remove_subcontext : forall x l l', remove x l l' -> x :: l' ⊆ l. 46 | Proof. 47 | intros; intros y ?. rewrite (In_remove _ _ _ _ H). 48 | destruct H0; auto. 49 | Qed. 50 | 51 | Lemma remove_subcontext_rev : forall x l l', remove x l l' -> l ⊆ x :: l'. 52 | Proof. 53 | intros. intros y ?. rewrite (In_remove _ _ _ _ H) in H0. 54 | simpl. destruct H0; auto. 55 | Qed. 56 | 57 | 58 | Inductive LJTstar_proves : list (prop atom) -> prop atom -> Prop := 59 | | LJTstar_init {Γ P} : In P Γ -> Γ ⇒* P 60 | | LJTstar_bot_elim {Γ P} : In ⊥ Γ -> Γ ⇒* P 61 | | LJTstar_top_intro {Γ} : Γ ⇒* ⊤ 62 | | LJTstar_and_intro {Γ P Q} : Γ ⇒* P -> Γ ⇒* Q -> Γ ⇒* P ∧ Q 63 | | LJTstar_and_elim {Γ P Q R Γ'} : remove (P ∧ Q) Γ Γ' -> 64 | P :: Q :: Γ' ⇒* R -> Γ ⇒* R 65 | | LJTstar_or_introl {Γ P Q} : Γ ⇒* P -> Γ ⇒* P ∨ Q 66 | | LJTstar_or_intror {Γ P Q} : Γ ⇒* Q -> Γ ⇒* P ∨ Q 67 | | LJTstar_or_elim {Γ P Q R Γ'} : remove (P ∨ Q) Γ Γ' -> 68 | P :: Γ' ⇒* R -> Q :: Γ' ⇒* R -> Γ ⇒* R 69 | | LJTstar_impl_intro {Γ P Q} : P :: Γ ⇒* Q -> Γ ⇒* P ⊃ Q 70 | | LJTstar_impl_assump_elim {Γ P Q R Γ'} : 71 | remove (P ⊃ Q) Γ Γ' -> In P Γ' -> Q :: Γ' ⇒* R -> 72 | Γ ⇒* R 73 | | LJTstar_red_top_impl {Γ P Q Γ'} : remove (⊤ ⊃ P) Γ Γ' -> 74 | P :: Γ' ⇒* Q -> Γ ⇒* Q 75 | | LJTstar_red_and_impl {Γ P Q R S Γ'} : remove (P ∧ Q ⊃ R) Γ Γ' -> 76 | P ⊃ Q ⊃ R :: Γ' ⇒* S -> Γ ⇒* S 77 | | LJTstar_red_or_impl {Γ P Q R S Γ'} : remove (P ∨ Q ⊃ R) Γ Γ' -> 78 | P ⊃ R :: Q ⊃ R :: Γ' ⇒* S -> Γ ⇒* S 79 | | LJTstar_impl_impl_elim {Γ P Q R S Γ'} : remove ((P ⊃ Q) ⊃ R) Γ Γ' -> 80 | P :: Q ⊃ R :: Γ' ⇒* Q -> R :: Γ' ⇒* S -> Γ ⇒* S 81 | where "Γ ⇒* P" := (LJTstar_proves Γ P). 82 | 83 | Example LJTstar_no_Peirce : forall x y:atom, x <> y -> let P := atom_prop x in 84 | let Q := atom_prop y in ~ (nil ⇒* ((P ⊃ Q) ⊃ P) ⊃ P). 85 | Proof. 86 | intros; subst P Q. intro. 87 | (* only possible first step is LJTstar_impl_intro *) 88 | inversion_clear H0; 89 | repeat match goal with 90 | | H0 : In _ _ |- _ => destruct H0; subst 91 | | H0 : remove _ _ _ |- _ => inversion_clear H0 92 | end. 93 | (* only possible next step is LJTstar_impl_impl_elim *) 94 | inversion H1; subst; 95 | repeat match goal with 96 | | H0 : In _ _ |- _ => destruct H0; try discriminate; subst 97 | | H0 : remove _ _ _ |- _ => inversion H0; subst 98 | end. 99 | (* in subgoal [ x; y ⊃ x ] ⇒* y, cannot make any more progress *) 100 | clear H1 H3 H0. inversion H2; subst; repeat match goal with 101 | | H0 : In _ _ |- _ => destruct H0; try discriminate; subst 102 | | H0 : remove _ _ _ |- _ => inversion H0; subst 103 | end; congruence. 104 | Qed. 105 | 106 | 107 | Theorem LJTstar_soundness {Γ P} : Γ ⇒* P -> Γ ⊢ P. 108 | Proof. 109 | induction 1; try match goal with 110 | | H : remove _ _ _ |- _ ⊢ _ => 111 | rewrite <- (remove_subcontext _ _ _ H); clear H 112 | end. 113 | + apply ND_assumption; assumption. 114 | + apply ND_exfalso_quodlibet. apply ND_assumption; assumption. 115 | + apply ND_True_intro. 116 | + apply ND_and_intro; assumption. 117 | + apply @ND_and_elim with (P := P) (Q := Q). 118 | - apply ND_assumption. prove_In. 119 | - refine (ND_context_extension _ _ _ _ _ eq_refl IHLJTstar_proves). 120 | prove_subcontext. 121 | + apply ND_or_introl; assumption. 122 | + apply ND_or_intror; assumption. 123 | + apply @ND_proof_by_cases with (P := P) (Q := Q). 124 | - apply ND_assumption. prove_In. 125 | - refine (ND_context_extension _ _ _ _ _ eq_refl IHLJTstar_proves1). 126 | prove_subcontext. 127 | - refine (ND_context_extension _ _ _ _ _ eq_refl IHLJTstar_proves2). 128 | prove_subcontext. 129 | + apply ND_cond_proof; assumption. 130 | + apply @ND_cut with (P := Q). 131 | - apply @ND_modus_ponens with (P := P). 132 | * apply ND_assumption. prove_In. 133 | * apply ND_assumption. right; assumption. 134 | - refine (ND_context_extension _ _ _ _ _ eq_refl IHLJTstar_proves). 135 | prove_subcontext. 136 | + apply @ND_cut with (P := P). 137 | - apply @ND_modus_ponens with (P := ⊤). 138 | * apply ND_assumption. prove_In. 139 | * apply ND_True_intro. 140 | - refine (ND_context_extension _ _ _ _ _ eq_refl IHLJTstar_proves). 141 | prove_subcontext. 142 | + apply @ND_cut with (P := P ⊃ Q ⊃ R). 143 | - do 2 apply ND_cond_proof. apply @ND_modus_ponens with (P := P ∧ Q). 144 | * apply ND_assumption. prove_In. 145 | * apply ND_and_intro; apply ND_assumption; prove_In. 146 | - refine (ND_context_extension _ _ _ _ _ eq_refl IHLJTstar_proves). 147 | prove_subcontext. 148 | + apply @ND_cut with (P := Q ⊃ R). 149 | - apply ND_cond_proof. apply @ND_modus_ponens with (P := P ∨ Q). 150 | * apply ND_assumption. prove_In. 151 | * apply ND_or_intror; apply ND_assumption; prove_In. 152 | - apply @ND_cut with (P := P ⊃ R). 153 | * apply ND_cond_proof. apply @ND_modus_ponens with (P := P ∨ Q). 154 | { apply ND_assumption. prove_In. } 155 | { apply ND_or_introl; apply ND_assumption; prove_In. } 156 | * refine (ND_context_extension _ _ _ _ _ eq_refl IHLJTstar_proves). 157 | prove_subcontext. 158 | + apply @ND_cut with (P := R). 159 | - apply @ND_modus_ponens with (P := P ⊃ Q). 160 | * apply ND_assumption; prove_In. 161 | * apply ND_cond_proof. apply @ND_cut with (P := Q ⊃ R). 162 | { apply ND_cond_proof. apply @ND_modus_ponens with (P := P ⊃ Q). 163 | + apply ND_assumption; prove_In. 164 | + apply ND_cond_proof. apply ND_assumption; prove_In. 165 | } 166 | { refine (ND_context_extension _ _ _ _ _ eq_refl IHLJTstar_proves1). 167 | prove_subcontext. } 168 | - refine (ND_context_extension _ _ _ _ _ eq_refl IHLJTstar_proves2). 169 | prove_subcontext. 170 | Qed. 171 | 172 | 173 | Fixpoint prop_cost (P : prop atom) : nat := 174 | match P with 175 | | atom_prop _ => 0 176 | | ⊤ => 0 177 | | ⊥ => 0 178 | | Q ∧ R => 2 + (prop_cost Q + prop_cost R) 179 | | Q ∨ R => 1 + (prop_cost Q + prop_cost R) 180 | | Q ⊃ R => 1 + (prop_cost Q + prop_cost R) 181 | end. 182 | 183 | Fixpoint add_cost_to_list (n:nat) (l:list nat) {struct n} : list nat := 184 | match n, l with 185 | | 0, nil => 1 :: nil 186 | | 0, h :: t => S h :: t 187 | | S m, nil => 0 :: add_cost_to_list m nil 188 | | S m, h :: t => h :: add_cost_to_list m t 189 | end. 190 | 191 | Fixpoint context_cost (Γ : list (prop atom)) : list nat := 192 | match Γ with 193 | | nil => nil 194 | | P :: Γ' => add_cost_to_list (prop_cost P) (context_cost Γ') 195 | end. 196 | 197 | Inductive context_cost_equiv : list nat -> list nat -> Prop := 198 | | context_cost_equiv_base : context_cost_equiv nil nil 199 | | context_cost_equiv_cons {h l r} : context_cost_equiv l r -> 200 | context_cost_equiv (h :: l) (h :: r) 201 | (* cases for unequal numbers of 0 at end *) 202 | | context_cost_equiv_nil_l {l} : context_cost_equiv l nil -> 203 | context_cost_equiv (0 :: l) nil 204 | | context_cost_equiv_nil_r {r} : context_cost_equiv nil r -> 205 | context_cost_equiv nil (0 :: r). 206 | 207 | Instance context_cost_equiv_equiv : Equivalence context_cost_equiv. 208 | Proof. 209 | constructor. 210 | + intro l; induction l; auto using context_cost_equiv. 211 | + intros l1 l2 H; induction H; auto using context_cost_equiv. 212 | + intros l1 l2 l3 H0 H1; revert l3 H1; induction H0; inversion 1; subst; 213 | auto using context_cost_equiv. 214 | Qed. 215 | 216 | Instance context_cost_cons_equiv_proper : 217 | Proper (eq ==> context_cost_equiv ==> context_cost_equiv) (@cons nat). 218 | Proof. 219 | intros m n [] l1 l2 ?; auto using context_cost_equiv_cons. 220 | Qed. 221 | 222 | (* reverse lexical ordering *) 223 | Inductive context_cost_lt : list nat -> list nat -> Prop := 224 | | context_cost_lt_later {h h' t t'} : 225 | context_cost_lt t t' -> context_cost_lt (h :: t) (h' :: t') 226 | | context_cost_lt_here {h h' t t'} : 227 | h < h' -> context_cost_equiv t t' -> context_cost_lt (h :: t) (h' :: t') 228 | | context_cost_lt_nil_here {h t} : 229 | 0 < h -> context_cost_lt nil (h :: t) 230 | | context_cost_lt_nil_later {h t} : 231 | context_cost_lt nil t -> context_cost_lt nil (h :: t). 232 | 233 | Instance context_cost_cons_lt_proper : 234 | Proper (eq ==> context_cost_lt ++> context_cost_lt) (@cons nat). 235 | Proof. 236 | intros m n [] l1 l2 ?; auto using context_cost_lt_later. 237 | Qed. 238 | 239 | Require Import Arith. 240 | 241 | Lemma context_cost_nil_le : 242 | forall l, context_cost_equiv nil l \/ context_cost_lt nil l. 243 | Proof. 244 | induction l; auto using context_cost_equiv, context_cost_lt. 245 | destruct IHl. 246 | + destruct a; auto using context_cost_equiv, context_cost_lt with arith. 247 | + right. auto using context_cost_lt. 248 | Qed. 249 | 250 | Instance context_cost_equiv_lt_proper : 251 | Proper (context_cost_equiv ==> context_cost_equiv ==> iff) context_cost_lt. 252 | Proof. 253 | cut (Proper (context_cost_equiv ==> context_cost_equiv ==> Basics.impl) 254 | context_cost_lt). 255 | + intros Himpl l l' H l0 l0' H0; split. 256 | - exact (Himpl l l' H l0 l0' H0). 257 | - symmetry in H, H0. exact (Himpl l' l H l0' l0 H0). 258 | + intros l l' H; induction H; intros l0 l0' H0; induction H0; 259 | unfold Basics.impl; inversion 1; subst; auto using context_cost_lt. 260 | all: try match goal with 261 | | H : _ < 0 |- _ => contradict H; auto with arith 262 | end. 263 | - apply context_cost_lt_later. eapply IHcontext_cost_equiv; eauto. 264 | - apply context_cost_lt_here; trivial. 265 | rewrite <- H. rewrite <- H0. assumption. 266 | - eapply IHcontext_cost_equiv in H3; eauto. inversion H3. 267 | - apply context_cost_lt_nil_later. eapply IHcontext_cost_equiv; eauto. 268 | - eapply IHcontext_cost_equiv in H3; eauto. 269 | - destruct (context_cost_nil_le r0). 270 | * apply context_cost_lt_here; trivial. 271 | rewrite <- H. assumption. 272 | * apply context_cost_lt_later. 273 | eapply IHcontext_cost_equiv in H2; eauto. reflexivity. 274 | - apply context_cost_lt_later. eapply IHcontext_cost_equiv; eauto. 275 | Qed. 276 | 277 | Instance context_cost_lt_trans : Transitive context_cost_lt. 278 | Proof. 279 | intros l1 l2 l3 H0 H1; revert l3 H1; induction H0; inversion 1; subst; 280 | auto using context_cost_lt. 281 | + apply context_cost_lt_later. rewrite <- H5. assumption. 282 | + apply context_cost_lt_later. rewrite H0. assumption. 283 | + apply context_cost_lt_here; eauto with arith. 284 | etransitivity; eauto. 285 | + apply context_cost_lt_nil_later. clear H1. 286 | induction H4; auto using context_cost_lt. 287 | apply context_cost_lt_nil_here. eauto with arith. 288 | + apply context_cost_lt_nil_here. eauto with arith. 289 | + apply context_cost_lt_nil_later. rewrite <- H5. assumption. 290 | Qed. 291 | 292 | Proposition context_cost_lt_wf : well_founded context_cost_lt. 293 | Proof. 294 | intro l; induction l. 295 | + constructor. inversion 1. 296 | + revert a; induction IHl; intro a. induction (lt_wf a). constructor. 297 | inversion 1; subst; eauto. 298 | - pose proof (H2 _ H7). constructor. intros. rewrite H8 in H5. 299 | exact (Acc_inv H4 H5). 300 | - constructor; inversion 1. 301 | Qed. 302 | 303 | Instance add_cost_to_list_proper : 304 | Proper (eq ==> context_cost_equiv ==> context_cost_equiv) 305 | add_cost_to_list. 306 | Proof. 307 | intros m n [] l1 l2 ?. clear n. revert l1 l2 H; induction m; intros. 308 | + destruct H; simpl. 309 | - reflexivity. 310 | - f_equiv. assumption. 311 | - f_equiv. assumption. 312 | - f_equiv. assumption. 313 | + destruct H; simpl. 314 | - reflexivity. 315 | - f_equiv; auto. 316 | - f_equiv; auto. 317 | - f_equiv; auto. 318 | Qed. 319 | 320 | Instance add_cost_to_list_incr : 321 | Proper (eq ==> context_cost_lt ++> context_cost_lt) 322 | add_cost_to_list. 323 | Proof. 324 | intros m n [] l1 l2 ?. clear n. revert l1 l2 H; induction m; intros. 325 | + destruct H; simpl; auto using context_cost_lt with arith. 326 | - destruct (context_cost_nil_le t); auto using context_cost_lt with arith. 327 | + destruct H; simpl; auto using context_cost_lt with arith. 328 | - apply context_cost_lt_here; trivial. repeat f_equiv; assumption. 329 | - destruct (context_cost_nil_le t); auto using context_cost_lt with arith. 330 | apply context_cost_lt_here; trivial. repeat f_equiv; assumption. 331 | Qed. 332 | 333 | Lemma add_cost_to_list_lt : 334 | forall n l, context_cost_lt l (add_cost_to_list n l). 335 | Proof. 336 | induction n; destruct l; simpl; auto using context_cost_lt with arith. 337 | apply context_cost_lt_here. 338 | + constructor. 339 | + reflexivity. 340 | Qed. 341 | 342 | Fixpoint nth_tail (n:nat) (l:list nat) : list nat := 343 | match n with 344 | | 0 => l 345 | | S m => match l with 346 | | nil => nil 347 | | _ :: l' => nth_tail m l' 348 | end 349 | end. 350 | 351 | Lemma nth_tail_nil : forall n, nth_tail n nil = nil. 352 | Proof. 353 | induction n; simpl; trivial. 354 | Qed. 355 | 356 | Lemma nth_tail_lt : forall n l l', 357 | context_cost_lt (nth_tail n l) (nth_tail n l') -> 358 | context_cost_lt l l'. 359 | Proof. 360 | induction n. 361 | + simpl; trivial. 362 | + destruct l, l'; simpl; auto using context_cost_lt. 363 | - intros. apply IHn. rewrite nth_tail_nil. clear IHn. revert n0 l' H. induction n. 364 | * intros. simpl in H |-*. auto using context_cost_lt. 365 | * intros. simpl. destruct l'; simpl in H; auto using context_cost_lt. 366 | inversion H. 367 | - inversion 1. 368 | Qed. 369 | 370 | Lemma nth_tail_add_cost_ge : forall n m l, 371 | m >= n -> context_cost_lt (nth_tail n l) (nth_tail n (add_cost_to_list m l)). 372 | Proof. 373 | induction n. 374 | + simpl. intros. apply add_cost_to_list_lt. 375 | + destruct l; intros; destruct m; try match goal with 376 | | H : 0 >= S _ |- _ => contradict H; unfold not; inversion 1 377 | end; simpl; auto with arith. 378 | rewrite <- (nth_tail_nil n) at 1. auto with arith. 379 | Qed. 380 | 381 | Lemma nth_tail_add_cost_lt : forall n m l, 382 | m < n -> nth_tail n l = nth_tail n (add_cost_to_list m l). 383 | Proof. 384 | intros. revert n l H; induction m; intros. 385 | + destruct n; try (inversion H; fail). destruct l; simpl. 386 | - rewrite nth_tail_nil; reflexivity. 387 | - reflexivity. 388 | + destruct n; try (inversion H; fail). destruct l; simpl. 389 | - rewrite <- (nth_tail_nil n) at 1. auto with arith. 390 | - auto with arith. 391 | Qed. 392 | 393 | Lemma add_cost_to_list_comm : forall n m l, 394 | add_cost_to_list n (add_cost_to_list m l) = 395 | add_cost_to_list m (add_cost_to_list n l). 396 | Proof. 397 | induction n; destruct m, l; simpl; f_equal; auto. 398 | Qed. 399 | 400 | Lemma remove_context_cost : forall Γ Γ' P, remove P Γ Γ' -> 401 | context_cost Γ = add_cost_to_list (prop_cost P) (context_cost Γ'). 402 | Proof. 403 | induction 1. 404 | + reflexivity. 405 | + simpl. rewrite IHremove. apply add_cost_to_list_comm. 406 | Qed. 407 | 408 | Proposition reduce_context_cost_lt : forall Γ Γ' Γ'' P, remove P Γ Γ' -> 409 | (forall Q, In Q Γ'' -> prop_cost Q < prop_cost P) -> 410 | context_cost_lt (context_cost (Γ'' ++ Γ')) (context_cost Γ). 411 | Proof. 412 | intros. rewrite (remove_context_cost _ _ _ H). 413 | apply nth_tail_lt with (n := prop_cost P). 414 | assert (context_cost_equiv (nth_tail (prop_cost P) (context_cost (Γ'' ++ Γ'))) 415 | (nth_tail (prop_cost P) (context_cost Γ'))). 416 | + clear Γ H. induction Γ''. 417 | - reflexivity. 418 | - simpl. rewrite <- nth_tail_add_cost_lt. 419 | * apply IHΓ''. intros. apply H0. right; assumption. 420 | * apply H0. prove_In. 421 | + rewrite H1. apply nth_tail_add_cost_ge. constructor. 422 | Qed. 423 | 424 | 425 | Instance add_cost_to_list_incr_l : 426 | Proper (lt ++> context_cost_equiv ==> context_cost_lt) add_cost_to_list. 427 | Proof. 428 | intros m n ? l1 l2 ?. rewrite H0. clear l1 H0. 429 | revert n H l2; induction m; destruct n, l2; intros; simpl; 430 | try match goal with 431 | | H : _ < 0 |- _ => inversion H 432 | end. 433 | + apply context_cost_lt_later. apply add_cost_to_list_lt. 434 | + apply context_cost_lt_later. apply add_cost_to_list_lt. 435 | + apply context_cost_lt_later. apply IHm. auto with arith. 436 | + apply context_cost_lt_later. apply IHm. auto with arith. 437 | Qed. 438 | 439 | (* Now, for a series of results that each subgoal in an LJT* rule 440 | reduces the cost *) 441 | Lemma context_cost_lt_cons_l {Γ P Q} : 442 | prop_cost P < prop_cost Q -> 443 | context_cost_lt (context_cost (P :: Γ)) (context_cost (Q :: Γ)). 444 | Proof. 445 | intros. simpl. apply add_cost_to_list_incr_l; trivial. reflexivity. 446 | Qed. 447 | 448 | Lemma context_cost_lt_cons_r {Γ Γ' P} : 449 | context_cost_lt (context_cost Γ) (context_cost Γ') -> 450 | context_cost_lt (context_cost (P :: Γ)) (context_cost (P :: Γ')). 451 | Proof. 452 | intros. simpl. apply add_cost_to_list_incr; trivial. 453 | Qed. 454 | 455 | Proposition LJTstar_and_intro_cost1 {Γ P Q} : 456 | context_cost_lt (context_cost (P :: Γ)) (context_cost (P ∧ Q :: Γ)). 457 | Proof. 458 | apply context_cost_lt_cons_l. simpl. eauto with arith. 459 | Qed. 460 | 461 | Proposition LJTstar_and_intro_cost2 {Γ P Q} : 462 | context_cost_lt (context_cost (Q :: Γ)) (context_cost (P ∧ Q :: Γ)). 463 | Proof. 464 | apply context_cost_lt_cons_l. simpl. eauto with arith. 465 | Qed. 466 | 467 | Proposition LJTstar_and_elim_cost {Γ P Q R Γ'} : 468 | remove (P ∧ Q) Γ Γ' -> 469 | context_cost_lt (context_cost (R :: P :: Q :: Γ')) 470 | (context_cost (R :: Γ)). 471 | Proof. 472 | intros. apply context_cost_lt_cons_r. 473 | apply reduce_context_cost_lt with (1 := H) (Γ'' := P :: Q :: nil). 474 | intros; repeat match goal with 475 | | H : In _ (_ :: _) |- _ => destruct H; subst 476 | | H : In _ nil |- _ => destruct H 477 | end; simpl; eauto with arith. 478 | Qed. 479 | 480 | Proposition LJTstar_or_introl_cost {Γ P Q} : 481 | context_cost_lt (context_cost (P :: Γ)) 482 | (context_cost (P ∨ Q :: Γ)). 483 | Proof. 484 | apply context_cost_lt_cons_l. simpl; eauto with arith. 485 | Qed. 486 | 487 | Proposition LJTstar_or_intror_cost {Γ P Q} : 488 | context_cost_lt (context_cost (Q :: Γ)) 489 | (context_cost (P ∨ Q :: Γ)). 490 | Proof. 491 | apply context_cost_lt_cons_l. simpl; eauto with arith. 492 | Qed. 493 | 494 | Proposition LJTstar_or_elim_cost1 {Γ P Q R Γ'} : 495 | remove (P ∨ Q) Γ Γ' -> context_cost_lt 496 | (context_cost (R :: P :: Γ')) 497 | (context_cost (R :: Γ)). 498 | Proof. 499 | intros. apply context_cost_lt_cons_r. 500 | apply reduce_context_cost_lt with (1 := H) (Γ'' := P :: nil). 501 | intros; repeat match goal with 502 | | H : In _ (_ :: _) |- _ => destruct H; subst 503 | | H : In _ nil |- _ => destruct H 504 | end; simpl; eauto with arith. 505 | Qed. 506 | 507 | Proposition LJTstar_or_elim_cost2 {Γ P Q R Γ'} : 508 | remove (P ∨ Q) Γ Γ' -> context_cost_lt 509 | (context_cost (R :: Q :: Γ')) 510 | (context_cost (R :: Γ)). 511 | Proof. 512 | intros. apply context_cost_lt_cons_r. 513 | apply reduce_context_cost_lt with (1 := H) (Γ'' := Q :: nil). 514 | intros; repeat match goal with 515 | | H : In _ (_ :: _) |- _ => destruct H; subst 516 | | H : In _ nil |- _ => destruct H 517 | end; simpl; eauto with arith. 518 | Qed. 519 | 520 | Proposition LJTstar_impl_intro_cost {Γ P Q} : 521 | context_cost_lt (context_cost (Q :: P :: Γ)) 522 | (context_cost (P ⊃ Q :: Γ)). 523 | Proof. 524 | assert (remove (P ⊃ Q) (P ⊃ Q :: Γ) Γ) by eauto using remove. 525 | apply reduce_context_cost_lt with (1 := H) (Γ'' := Q :: P :: nil). 526 | intros; repeat match goal with 527 | | H : In _ (_ :: _) |- _ => destruct H; subst 528 | | H : In _ nil |- _ => destruct H 529 | end; simpl; eauto with arith. 530 | Qed. 531 | 532 | Proposition LJTstar_impl_assump_elim_cost {Γ P Q R Γ'} : 533 | remove (P ⊃ Q) Γ Γ' -> context_cost_lt 534 | (context_cost (R :: Q :: Γ')) 535 | (context_cost (R :: Γ)). 536 | Proof. 537 | intros. apply context_cost_lt_cons_r. 538 | apply reduce_context_cost_lt with (1 := H) (Γ'' := Q :: nil). 539 | intros; repeat match goal with 540 | | H : In _ (_ :: _) |- _ => destruct H; subst 541 | | H : In _ nil |- _ => destruct H 542 | end; simpl; eauto with arith. 543 | Qed. 544 | 545 | Proposition LJTstar_red_top_impl_cost {Γ P Q Γ'} : 546 | remove (⊤ ⊃ P) Γ Γ' -> context_cost_lt 547 | (context_cost (Q :: P :: Γ')) 548 | (context_cost (Q :: Γ)). 549 | Proof. 550 | intros. apply context_cost_lt_cons_r. 551 | apply reduce_context_cost_lt with (1 := H) (Γ'' := P :: nil). 552 | intros; repeat match goal with 553 | | H : In _ (_ :: _) |- _ => destruct H; subst 554 | | H : In _ nil |- _ => destruct H 555 | end; simpl; eauto with arith. 556 | Qed. 557 | 558 | Proposition LJTstar_red_and_impl_cost {Γ P Q R S Γ'} : 559 | remove (P ∧ Q ⊃ R) Γ Γ' -> context_cost_lt 560 | (context_cost (S :: P ⊃ Q ⊃ R :: Γ')) 561 | (context_cost (S :: Γ)). 562 | Proof. 563 | intros. apply context_cost_lt_cons_r. 564 | apply reduce_context_cost_lt with (1 := H) (Γ'' := P ⊃ Q ⊃ R :: nil). 565 | intros; repeat match goal with 566 | | H : In _ (_ :: _) |- _ => destruct H; subst 567 | | H : In _ nil |- _ => destruct H 568 | end; simpl. repeat rewrite <- plus_n_Sm. rewrite plus_assoc. eauto with arith. 569 | Qed. 570 | 571 | Proposition LJTstar_red_or_impl_cost {Γ P Q R S Γ'} : 572 | remove (P ∨ Q ⊃ R) Γ Γ' -> context_cost_lt 573 | (context_cost (S :: P ⊃ R :: Q ⊃ R :: Γ')) 574 | (context_cost (S :: Γ)). 575 | Proof. 576 | intros. apply context_cost_lt_cons_r. 577 | apply reduce_context_cost_lt with (1 := H) (Γ'' := P ⊃ R :: Q ⊃ R :: nil). 578 | intros; repeat match goal with 579 | | H : In _ (_ :: _) |- _ => destruct H; subst 580 | | H : In _ nil |- _ => destruct H 581 | end; simpl; eauto with arith. 582 | Qed. 583 | 584 | Proposition LJTstar_impl_impl_elim_cost1 {Γ P Q R S Γ'} : 585 | remove ((P ⊃ Q) ⊃ R) Γ Γ' -> context_cost_lt 586 | (context_cost (Q :: P :: Q ⊃ R :: Γ')) 587 | (context_cost (S :: Γ)). 588 | Proof. 589 | intros. simpl context_cost at 2. rewrite <- add_cost_to_list_lt. 590 | apply reduce_context_cost_lt with (1 := H) (Γ'' := Q :: P :: Q ⊃ R :: nil). 591 | intros; repeat match goal with 592 | | H : In _ (_ :: _) |- _ => destruct H; subst 593 | | H : In _ nil |- _ => destruct H 594 | end; simpl; eauto with arith. 595 | Qed. 596 | 597 | Proposition LJTstar_impl_impl_elim_cost2 {Γ P Q R S Γ'} : 598 | remove ((P ⊃ Q) ⊃ R) Γ Γ' -> context_cost_lt 599 | (context_cost (S :: R :: Γ')) 600 | (context_cost (S :: Γ)). 601 | Proof. 602 | intros. apply context_cost_lt_cons_r. 603 | apply reduce_context_cost_lt with (1 := H) (Γ'' := R :: nil). 604 | intros; repeat match goal with 605 | | H : In _ (_ :: _) |- _ => destruct H; subst 606 | | H : In _ nil |- _ => destruct H 607 | end; simpl; eauto with arith. 608 | Qed. 609 | 610 | 611 | Theorem LJTstar_completeness : 612 | forall Γ P, Γ ⊢ P -> Γ ⇒* P. 613 | Proof. 614 | intros. rewrite ND_SC_equiv in H. 615 | pose proof (context_cost_lt_wf (context_cost (P :: Γ))). 616 | remember (context_cost (P :: Γ)) as c. revert Γ P H Heqc. 617 | induction H0. intros. subst. 618 | pose proof (fun Γ' P' H2 H3 => H0 (context_cost (P' :: Γ')) H2 Γ' P' H3 eq_refl). 619 | clear H H0. 620 | destruct H1. 621 | + apply LJTstar_init. assumption. 622 | + apply LJTstar_bot_elim. assumption. 623 | + apply LJTstar_top_intro. 624 | + apply LJTstar_and_intro. 625 | - apply H2; trivial. apply LJTstar_and_intro_cost1. 626 | - apply H2; trivial. apply LJTstar_and_intro_cost2. 627 | + rewrite In_ex_remove in H. destruct H as [Γ']. 628 | apply (LJTstar_and_elim H). apply H2. 629 | - eauto using LJTstar_and_elim_cost. 630 | - apply @SC_admits_cut with (P := P ∧ Q). 631 | * apply SC_and_intro; apply SC_init; prove_In. 632 | * rewrite (remove_subcontext_rev _ _ _ H) in H1. 633 | refine (SC_context_extension _ _ _ _ _ eq_refl H1). 634 | prove_subcontext. 635 | + apply LJTstar_or_introl. apply H2; trivial. apply LJTstar_or_introl_cost. 636 | + apply LJTstar_or_intror. apply H2; trivial. apply LJTstar_or_intror_cost. 637 | + rewrite In_ex_remove in H. destruct H as [Γ']. 638 | apply (LJTstar_or_elim H); apply H2. 639 | - eauto using LJTstar_or_elim_cost1. 640 | - apply @SC_admits_cut with (P := P ∨ Q). 641 | * apply SC_or_introl; apply SC_init; prove_In. 642 | * rewrite (remove_subcontext_rev _ _ _ H) in H1_. 643 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_). 644 | prove_subcontext. 645 | - eauto using LJTstar_or_elim_cost2. 646 | - apply @SC_admits_cut with (P := P ∨ Q). 647 | * apply SC_or_intror; apply SC_init; prove_In. 648 | * rewrite (remove_subcontext_rev _ _ _ H) in H1_0. 649 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_0). 650 | prove_subcontext. 651 | + apply LJTstar_impl_intro. apply H2; trivial. 652 | apply LJTstar_impl_intro_cost. 653 | + revert Q R H H1_0 H2; induction H1_; intros Q0 R0 H' H1_0' H2; 654 | rewrite In_ex_remove in H'; destruct H' as [Γ']. 655 | - apply (LJTstar_impl_assump_elim H0). 656 | * pose proof (remove_subcontext_rev _ _ _ H0). apply H1 in H. 657 | destruct H; trivial. contradict H. 658 | clear H0 H1_0' H1; revert Q0; induction P; try discriminate. 659 | unfold not; injection 1; intros. contradiction (IHP1 P2). 660 | * apply H2. 661 | { eauto using LJTstar_impl_assump_elim_cost. } 662 | { apply @SC_admits_cut with (P := P ⊃ Q0). 663 | + apply SC_impl_intro. apply SC_init; prove_In. 664 | + rewrite (remove_subcontext_rev _ _ _ H0) in H1_0'. 665 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_0'). 666 | prove_subcontext. } 667 | - apply LJTstar_bot_elim; assumption. 668 | - apply (LJTstar_red_top_impl H). apply H2. 669 | * eauto using LJTstar_red_top_impl_cost. 670 | * apply @SC_admits_cut with (P := ⊤ ⊃ Q0). 671 | { apply SC_impl_intro; apply SC_init; prove_In. } 672 | { rewrite (remove_subcontext_rev _ _ _ H) in H1_0'. 673 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_0'). 674 | prove_subcontext. } 675 | - apply (LJTstar_red_and_impl H). apply H2. 676 | * eauto using LJTstar_red_and_impl_cost. 677 | * apply @SC_admits_cut with (P := P ∧ Q ⊃ Q0). 678 | { apply SC_impl_intro. 679 | apply @SC_and_elim with (P := P) (Q := Q); try prove_In. 680 | apply @SC_impl_elim with (P := P) (Q := Q ⊃ Q0); try prove_In. 681 | + apply SC_init; prove_In. 682 | + apply @SC_impl_elim with (P := Q) (Q := Q0); try prove_In. 683 | - apply SC_init; prove_In. 684 | - apply SC_init; prove_In. } 685 | { assert (Γ ⇒ R0). 686 | + apply @SC_impl_elim with (P := P ∧ Q) (Q := Q0). 687 | - apply (remove_subcontext _ _ _ H). prove_In. 688 | - apply SC_and_intro; assumption. 689 | - assumption. 690 | + rewrite (remove_subcontext_rev _ _ _ H) in H0. 691 | refine (SC_context_extension _ _ _ _ _ eq_refl H0). 692 | prove_subcontext. 693 | } 694 | - rewrite In_ex_remove in H. destruct H as [Γ'']. 695 | apply (LJTstar_and_elim H). apply H2. 696 | * eauto using LJTstar_and_elim_cost. 697 | * apply @SC_admits_cut with (P := R). 698 | { apply @SC_admits_cut with (P := P ∧ Q). 699 | + apply SC_and_intro; apply SC_init; prove_In. 700 | + rewrite (remove_subcontext_rev _ _ _ H) in H1_. 701 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_). 702 | prove_subcontext. } 703 | { apply @SC_admits_cut with (P := Q0). 704 | + apply @SC_impl_elim with (P := R) (Q := Q0). 705 | - apply remove_In in H0. apply remove_subcontext_rev in H. 706 | apply H in H0. destruct H0; try discriminate. 707 | do 3 right; assumption. 708 | - apply SC_init; prove_In. 709 | - apply SC_init; prove_In. 710 | + apply @SC_admits_cut with (P := P ∧ Q). 711 | - apply SC_and_intro; apply SC_init; prove_In. 712 | - rewrite (remove_subcontext_rev _ _ _ H) in H1_0'. 713 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_0'). 714 | prove_subcontext. } 715 | - apply (LJTstar_red_or_impl H). apply H2. 716 | * eauto using LJTstar_red_or_impl_cost. 717 | * apply @SC_admits_cut with (P := P ∨ Q ⊃ Q0). 718 | { apply SC_impl_intro. 719 | apply @SC_or_elim with (P := P) (Q := Q); try prove_In. 720 | + apply @SC_impl_elim with (P := P) (Q := Q0); try prove_In. 721 | - apply SC_init; prove_In. 722 | - apply SC_init; prove_In. 723 | + apply @SC_impl_elim with (P := Q) (Q := Q0); try prove_In. 724 | - apply SC_init; prove_In. 725 | - apply SC_init; prove_In. } 726 | { assert (Γ ⇒ R0). 727 | + apply @SC_impl_elim with (P := P ∨ Q) (Q := Q0). 728 | - apply (remove_subcontext _ _ _ H). prove_In. 729 | - apply SC_or_introl; assumption. 730 | - assumption. 731 | + rewrite (remove_subcontext_rev _ _ _ H) in H0. 732 | refine (SC_context_extension _ _ _ _ _ eq_refl H0). 733 | prove_subcontext. } 734 | - apply (LJTstar_red_or_impl H). apply H2. 735 | * eauto using LJTstar_red_or_impl_cost. 736 | * apply @SC_admits_cut with (P := P ∨ Q ⊃ Q0). 737 | { apply SC_impl_intro. 738 | apply @SC_or_elim with (P := P) (Q := Q); try prove_In. 739 | + apply @SC_impl_elim with (P := P) (Q := Q0); try prove_In. 740 | - apply SC_init; prove_In. 741 | - apply SC_init; prove_In. 742 | + apply @SC_impl_elim with (P := Q) (Q := Q0); try prove_In. 743 | - apply SC_init; prove_In. 744 | - apply SC_init; prove_In. } 745 | { assert (Γ ⇒ R0). 746 | + apply @SC_impl_elim with (P := P ∨ Q) (Q := Q0). 747 | - apply (remove_subcontext _ _ _ H). prove_In. 748 | - apply SC_or_intror; assumption. 749 | - assumption. 750 | + rewrite (remove_subcontext_rev _ _ _ H) in H0. 751 | refine (SC_context_extension _ _ _ _ _ eq_refl H0). 752 | prove_subcontext. } 753 | - rewrite In_ex_remove in H. destruct H as [Γ'']. 754 | apply (LJTstar_or_elim H); apply H2. 755 | * eauto using LJTstar_or_elim_cost1. 756 | * apply @SC_admits_cut with (P := R). 757 | { apply @SC_admits_cut with (P := P ∨ Q). 758 | + apply SC_or_introl; apply SC_init; prove_In. 759 | + rewrite (remove_subcontext_rev _ _ _ H) in H1_1. 760 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_1). 761 | prove_subcontext. } 762 | { apply @SC_impl_elim with (P := R) (Q := Q0). 763 | + right; right. apply remove_In in H0. 764 | apply (remove_subcontext_rev _ _ _ H) in H0. 765 | destruct H0; try discriminate. assumption. 766 | + apply SC_init; prove_In. 767 | + apply @SC_admits_cut with (P := P ∨ Q). 768 | - apply SC_or_introl. apply SC_init; prove_In. 769 | - rewrite (remove_subcontext_rev _ _ _ H) in H1_0'. 770 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_0'). 771 | prove_subcontext. } 772 | * eauto using LJTstar_or_elim_cost2. 773 | * apply @SC_admits_cut with (P := R). 774 | { apply @SC_admits_cut with (P := P ∨ Q). 775 | + apply SC_or_intror; apply SC_init; prove_In. 776 | + rewrite (remove_subcontext_rev _ _ _ H) in H1_2. 777 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_2). 778 | prove_subcontext. } 779 | { apply @SC_impl_elim with (P := R) (Q := Q0). 780 | + right; right. apply remove_In in H0. 781 | apply (remove_subcontext_rev _ _ _ H) in H0. 782 | destruct H0; try discriminate. assumption. 783 | + apply SC_init; prove_In. 784 | + apply @SC_admits_cut with (P := P ∨ Q). 785 | - apply SC_or_intror. apply SC_init; prove_In. 786 | - rewrite (remove_subcontext_rev _ _ _ H) in H1_0'. 787 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_0'). 788 | prove_subcontext. } 789 | - apply (LJTstar_impl_impl_elim H). 790 | * apply H2. 791 | { eauto using LJTstar_impl_impl_elim_cost1. } 792 | { apply @SC_admits_cut with (P := (P ⊃ Q) ⊃ Q0). 793 | + apply SC_impl_intro. 794 | apply @SC_impl_elim with (P := P) (Q := Q); 795 | try (prove_In || apply SC_init; prove_In). 796 | apply @SC_impl_elim with (P := Q) (Q := Q0); 797 | (prove_In || apply SC_init; prove_In). 798 | + rewrite (remove_subcontext_rev _ _ _ H) in H1_. 799 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_). 800 | prove_subcontext. } 801 | * apply H2. 802 | { eauto using LJTstar_impl_impl_elim_cost2. } 803 | { apply @SC_admits_cut with (P := (P ⊃ Q) ⊃ Q0). 804 | + apply SC_impl_intro. apply SC_init; prove_In. 805 | + rewrite (remove_subcontext_rev _ _ _ H) in H1_0'. 806 | refine (SC_context_extension _ _ _ _ _ eq_refl H1_0'). 807 | prove_subcontext. } 808 | - eapply IHH1_1; eauto. apply @SC_admits_cut with (P := R). 809 | * assumption. 810 | * pose proof (remove_In _ _ _ H0). 811 | apply @SC_impl_elim with (P := R) (Q := Q0). 812 | { right; right; assumption. } 813 | { apply SC_init; prove_In. } 814 | { refine (SC_context_extension _ _ _ _ _ eq_refl H1_0'). 815 | prove_subcontext. } 816 | Qed. 817 | 818 | Theorem ND_LJTstar_equiv : forall Γ P, 819 | Γ ⊢ P <-> Γ ⇒* P. 820 | Proof. 821 | split; [ apply LJTstar_completeness | apply LJTstar_soundness ]. 822 | Qed. 823 | 824 | End LJTstar. 825 | -------------------------------------------------------------------------------- /NDSCEquiv.v: -------------------------------------------------------------------------------- 1 | Require Import NaturalDeduction. 2 | Require Import SequentCalculus. 3 | Require Import List. 4 | Require Import Subcontext. 5 | 6 | Lemma ND_SC_equiv {atom:Type} : 7 | forall (Γ : list (prop atom)) (P : prop atom), 8 | Γ ⊢ P <-> Γ ⇒ P. 9 | Proof. 10 | split; intros. 11 | + induction H. 12 | - apply SC_admits_cut with (1 := IHND_proves). 13 | apply SC_bot_elim. left; reflexivity. 14 | - apply SC_top_intro. 15 | - apply SC_or_introl; trivial. 16 | - apply SC_or_intror; trivial. 17 | - apply SC_admits_cut with (1 := IHND_proves1). 18 | apply @SC_or_elim with (P := P) (Q := Q). 19 | * left; reflexivity. 20 | * refine (SC_context_extension _ _ _ _ _ eq_refl IHND_proves2). 21 | prove_subcontext. 22 | * refine (SC_context_extension _ _ _ _ _ eq_refl IHND_proves3). 23 | prove_subcontext. 24 | - apply SC_and_intro; assumption. 25 | - apply SC_admits_cut with (1 := IHND_proves1). 26 | apply @SC_and_elim with (P := P) (Q := Q). 27 | * left; reflexivity. 28 | * refine (SC_context_extension _ _ _ _ _ eq_refl IHND_proves2). 29 | prove_subcontext. 30 | - apply SC_impl_intro. assumption. 31 | - apply SC_admits_cut with (1 := IHND_proves1). 32 | apply @SC_impl_elim with (P := P) (Q := Q). 33 | * left; reflexivity. 34 | * refine (SC_context_extension _ _ _ _ _ eq_refl IHND_proves2). 35 | apply subcontext_cons_r. 36 | * apply SC_init. left; reflexivity. 37 | - apply SC_init; assumption. 38 | - eapply SC_admits_cut; eassumption. 39 | + induction H. 40 | - apply ND_assumption; trivial. 41 | - apply ND_exfalso_quodlibet. apply ND_assumption; trivial. 42 | - apply ND_True_intro. 43 | - apply ND_and_intro; trivial. 44 | - apply @ND_and_elim with (P := P) (Q := Q). 45 | * apply ND_assumption; trivial. 46 | * assumption. 47 | - apply ND_or_introl; trivial. 48 | - apply ND_or_intror; trivial. 49 | - apply @ND_proof_by_cases with (P := P) (Q := Q). 50 | * apply ND_assumption; trivial. 51 | * assumption. 52 | * assumption. 53 | - apply ND_cond_proof. assumption. 54 | - apply @ND_cut with (P := Q). 55 | * apply @ND_modus_ponens with (P := P). 56 | { apply ND_assumption; trivial. } 57 | { assumption. } 58 | * assumption. 59 | Qed. 60 | 61 | Corollary ND_consistent {atom : Type} : ~ (nil ⊢ @bot_prop atom). 62 | Proof. 63 | rewrite ND_SC_equiv. apply SC_consistent. 64 | Qed. 65 | -------------------------------------------------------------------------------- /NaturalDeduction.v: -------------------------------------------------------------------------------- 1 | Require Export PropLang. 2 | Require Export List. 3 | Require Export Morphisms. 4 | Require Export Subcontext. 5 | 6 | Section natural_deduction. 7 | 8 | Context {atom : Type}. 9 | 10 | Reserved Notation "Γ ⊢ P" (no associativity, at level 61). 11 | 12 | Inductive ND_proves : list (prop atom) -> prop atom -> Prop := 13 | | ND_exfalso_quodlibet {Γ P} : 14 | Γ ⊢ ⊥ -> 15 | Γ ⊢ P 16 | | ND_True_intro {Γ} : 17 | Γ ⊢ ⊤ 18 | | ND_or_introl {Γ P Q} : 19 | Γ ⊢ P -> 20 | Γ ⊢ P ∨ Q 21 | | ND_or_intror {Γ P Q} : 22 | Γ ⊢ Q -> 23 | Γ ⊢ P ∨ Q 24 | | ND_proof_by_cases {Γ P Q R} : 25 | Γ ⊢ P ∨ Q -> 26 | P :: Γ ⊢ R -> 27 | Q :: Γ ⊢ R -> 28 | Γ ⊢ R 29 | | ND_and_intro {Γ P Q} : 30 | Γ ⊢ P -> 31 | Γ ⊢ Q -> 32 | Γ ⊢ P ∧ Q 33 | | ND_and_elim {Γ P Q R} : 34 | Γ ⊢ P ∧ Q -> 35 | P :: Q :: Γ ⊢ R -> 36 | Γ ⊢ R 37 | | ND_cond_proof {Γ P Q} : 38 | P :: Γ ⊢ Q -> 39 | Γ ⊢ P ⊃ Q 40 | | ND_modus_ponens {Γ P Q} : 41 | Γ ⊢ P ⊃ Q -> 42 | Γ ⊢ P -> 43 | Γ ⊢ Q 44 | | ND_assumption {Γ P} : 45 | In P Γ -> 46 | Γ ⊢ P 47 | | ND_cut {Γ P Q} : 48 | Γ ⊢ P -> 49 | P :: Γ ⊢ Q -> 50 | Γ ⊢ Q 51 | where "Γ ⊢ P" := (ND_proves Γ P). 52 | 53 | Global Instance ND_context_extension : 54 | Proper (subcontext ++> eq ==> Basics.impl) ND_proves. 55 | Proof. 56 | intros Γ₁ Γ₂ ? P Q [] ?. revert Γ₂ H. induction H0; intros. 57 | + apply ND_exfalso_quodlibet. auto. 58 | + apply ND_True_intro. 59 | + apply ND_or_introl. auto. 60 | + apply ND_or_intror. auto. 61 | + apply (ND_proof_by_cases (P := P) (Q := Q0)); auto. 62 | - apply IHND_proves2. f_equiv. assumption. 63 | - apply IHND_proves3. f_equiv. assumption. 64 | + apply ND_and_intro; auto. 65 | + apply (ND_and_elim (P := P) (Q := Q0)); auto. 66 | apply IHND_proves2. do 2 f_equiv; assumption. 67 | + apply ND_cond_proof. apply IHND_proves. f_equiv; assumption. 68 | + apply (ND_modus_ponens (P := P)); auto. 69 | + apply ND_assumption. auto. 70 | + apply (ND_cut (P := P)); auto. 71 | apply IHND_proves2. f_equiv. assumption. 72 | Qed. 73 | 74 | (* Want to prove: ND_prop forms a Heyting algebra *) 75 | Definition ND_prop_le (P Q:prop atom) : Prop := 76 | ND_proves (cons P nil) Q. 77 | 78 | Definition ND_prop_eqv (P Q:prop atom) : Prop := 79 | ND_prop_le P Q /\ ND_prop_le Q P. 80 | 81 | Lemma ND_prop_le_refl {P} : ND_prop_le P P. 82 | Proof. 83 | apply ND_assumption. left. reflexivity. 84 | Qed. 85 | 86 | Lemma ND_prop_le_trans {P Q R} : 87 | ND_prop_le P Q -> ND_prop_le Q R -> ND_prop_le P R. 88 | Proof. 89 | intros. apply (ND_cut (P := Q)). 90 | + assumption. 91 | + refine (ND_context_extension _ _ _ _ _ eq_refl H0). prove_subcontext. 92 | Qed. 93 | 94 | Lemma ND_meet1 {P Q} : ND_prop_le (P ∧ Q) P. 95 | Proof. 96 | apply (ND_and_elim (P := P) (Q := Q)). 97 | + apply ND_assumption. left. reflexivity. 98 | + apply ND_assumption. left. reflexivity. 99 | Qed. 100 | 101 | Lemma ND_meet2 {P Q} : ND_prop_le (P ∧ Q) Q. 102 | Proof. 103 | apply (ND_and_elim (P := P) (Q := Q)). 104 | + apply ND_assumption. left. reflexivity. 105 | + apply ND_assumption. right. left. reflexivity. 106 | Qed. 107 | 108 | Lemma ND_meet_max {P Q R} : ND_prop_le R P -> ND_prop_le R Q -> 109 | ND_prop_le R (P ∧ Q). 110 | Proof. 111 | intros. apply ND_and_intro. 112 | + assumption. 113 | + assumption. 114 | Qed. 115 | 116 | Lemma ND_join1 {P Q} : ND_prop_le P (P ∨ Q). 117 | Proof. 118 | apply ND_or_introl. apply ND_assumption. left. reflexivity. 119 | Qed. 120 | 121 | Lemma ND_join2 {P Q} : ND_prop_le Q (P ∨ Q). 122 | Proof. 123 | apply ND_or_intror. apply ND_assumption. left. reflexivity. 124 | Qed. 125 | 126 | Lemma ND_join_min {P Q R} : ND_prop_le P R -> ND_prop_le Q R -> 127 | ND_prop_le (P ∨ Q) R. 128 | Proof. 129 | intros. apply (ND_proof_by_cases (P := P) (Q := Q)). 130 | + apply ND_assumption. left. reflexivity. 131 | + refine (ND_context_extension _ _ _ _ _ eq_refl H). prove_subcontext. 132 | + refine (ND_context_extension _ _ _ _ _ eq_refl H0). prove_subcontext. 133 | Qed. 134 | 135 | Lemma ND_False_min {P} : ND_prop_le ⊥ P. 136 | Proof. 137 | apply ND_exfalso_quodlibet. apply ND_assumption. 138 | left. reflexivity. 139 | Qed. 140 | 141 | Lemma ND_True_max {P} : ND_prop_le P ⊤. 142 | Proof. 143 | apply ND_True_intro. 144 | Qed. 145 | 146 | Lemma ND_Heyting_algebra {P Q R} : 147 | ND_prop_le P (Q ⊃ R) <-> ND_prop_le (P ∧ Q) R. 148 | Proof. 149 | split; intros. 150 | + apply (ND_and_elim (P := P) (Q := Q)). 151 | - apply ND_assumption. left. reflexivity. 152 | - apply (ND_modus_ponens (P := Q)). 153 | * refine (ND_context_extension _ _ _ _ _ eq_refl H). 154 | prove_subcontext. 155 | * apply ND_assumption. prove_In. 156 | + apply ND_cond_proof. apply (ND_cut (P := P ∧ Q)). 157 | - apply ND_and_intro. 158 | * apply ND_assumption. right. left. reflexivity. 159 | * apply ND_assumption. left. reflexivity. 160 | - refine (ND_context_extension _ _ _ _ _ eq_refl H). 161 | prove_subcontext. 162 | Qed. 163 | 164 | Section ND_free_Heyting_algebra. 165 | 166 | Variable X : Type. 167 | Variable (le : X -> X -> Prop). 168 | Variables (meet join arrow : X -> X -> X) (top bot : X). 169 | Variable f : atom -> X. 170 | 171 | Hypotheses 172 | (le_refl : forall x:X, le x x) 173 | (le_trans : forall x y z:X, le x y -> le y z -> le x z) 174 | (bot_min : forall x:X, le bot x) 175 | (top_max : forall x:X, le x top) 176 | (meet_left : forall x y:X, le (meet x y) x) 177 | (meet_right : forall x y:X, le (meet x y) y) 178 | (meet_max : forall x y z:X, le x y -> le x z -> le x (meet y z)) 179 | (join_left : forall x y:X, le x (join x y)) 180 | (join_right : forall x y:X, le y (join x y)) 181 | (join_min : forall x y z:X, le x z -> le y z -> le (join x y) z) 182 | (Heyting_cond : forall x y z:X, le x (arrow y z) <-> le (meet x y) z). 183 | 184 | Fixpoint F (P : prop atom) : X := 185 | match P with 186 | | atom_prop a => f a 187 | | ⊥ => bot 188 | | ⊤ => top 189 | | Q ∨ R => join (F Q) (F R) 190 | | Q ∧ R => meet (F Q) (F R) 191 | | Q ⊃ R => arrow (F Q) (F R) 192 | end. 193 | 194 | Lemma proves_cond {Γ P} : Γ ⊢ P -> 195 | le (F (fold_right and_prop ⊤ Γ)) (F P). 196 | Proof. 197 | induction 1; simpl in * |-*. 198 | + apply le_trans with (1 := IHND_proves). apply bot_min. 199 | + apply top_max. 200 | + apply le_trans with (1 := IHND_proves). apply join_left. 201 | + apply le_trans with (1 := IHND_proves). apply join_right. 202 | + rewrite <- Heyting_cond in IHND_proves2, IHND_proves3. 203 | pose proof (join_min _ _ _ IHND_proves2 IHND_proves3). 204 | rewrite Heyting_cond in H2. apply le_trans with (2 := H2). 205 | apply meet_max; auto. 206 | + apply meet_max; auto. 207 | + apply le_trans with (2 := IHND_proves2). apply meet_max. 208 | - apply le_trans with (1 := IHND_proves1). apply meet_left. 209 | - apply meet_max; auto. apply le_trans with (1 := IHND_proves1). 210 | apply meet_right. 211 | + rewrite Heyting_cond. apply le_trans with (2 := IHND_proves). auto. 212 | + rewrite Heyting_cond in IHND_proves1. rewrite <- Heyting_cond in IHND_proves1. 213 | pose proof (meet_max _ _ _ IHND_proves1 IHND_proves2). 214 | apply le_trans with (1 := H1). rewrite <- Heyting_cond. apply le_refl. 215 | + induction Γ; simpl. 216 | - destruct H. 217 | - destruct H. 218 | * rewrite <- H. apply meet_left. 219 | * apply le_trans with (2 := IHΓ H). apply meet_right. 220 | + apply le_trans with (2 := IHND_proves2). apply meet_max. 221 | - assumption. 222 | - apply le_refl. 223 | Qed. 224 | 225 | Corollary F_incr : forall P Q, ND_prop_le P Q -> le (F P) (F Q). 226 | Proof. 227 | intros. apply le_trans with (2 := proves_cond H). simpl. 228 | apply meet_max. 229 | + apply le_refl. 230 | + apply top_max. 231 | Qed. 232 | 233 | End ND_free_Heyting_algebra. 234 | 235 | End natural_deduction. 236 | 237 | Infix "⊢" := ND_proves (no associativity, at level 61). 238 | -------------------------------------------------------------------------------- /PropLang.v: -------------------------------------------------------------------------------- 1 | Section PropositionLanguage. 2 | 3 | Context { atom : Type }. 4 | 5 | Inductive prop : Type := 6 | | atom_prop : atom -> prop 7 | | bot_prop : prop 8 | | top_prop : prop 9 | | and_prop : prop -> prop -> prop 10 | | or_prop : prop -> prop -> prop 11 | | impl_prop : prop -> prop -> prop. 12 | 13 | Definition not_prop (P : prop) := 14 | impl_prop P bot_prop. 15 | 16 | End PropositionLanguage. 17 | 18 | Arguments prop atom : clear implicits. 19 | 20 | Notation "⊥" := bot_prop. 21 | Notation "⊤" := top_prop. 22 | Notation "¬ P" := (not_prop P) (at level 51). 23 | Infix "∧" := and_prop (left associativity, at level 52). 24 | Infix "∨" := or_prop (left associativity, at level 53). 25 | Infix "⊃" := impl_prop (right associativity, at level 54). 26 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # coq-sequent-calculus 2 | Coq formalizations of Sequent Calculus, Natural Deduction, etc. systems for propositional logic 3 | -------------------------------------------------------------------------------- /SequentCalculus.v: -------------------------------------------------------------------------------- 1 | Require Export PropLang. 2 | Require Export List. 3 | Require Export Subcontext. 4 | 5 | Section SequentCalculus. 6 | 7 | Context {atom : Type}. 8 | 9 | Reserved Notation "Γ ⇒ P" (no associativity, at level 61). 10 | 11 | Inductive SC_proves : list (prop atom) -> prop atom -> Prop := 12 | | SC_init {P Γ} : In P Γ -> Γ ⇒ P 13 | | SC_bot_elim {P Γ} : In ⊥ Γ -> Γ ⇒ P 14 | | SC_top_intro {Γ} : Γ ⇒ ⊤ 15 | | SC_and_intro {Γ P Q} : Γ ⇒ P -> Γ ⇒ Q -> Γ ⇒ P ∧ Q 16 | | SC_and_elim {Γ P Q R} : In (P ∧ Q) Γ -> P :: Q :: Γ ⇒ R -> Γ ⇒ R 17 | | SC_or_introl {Γ P Q} : Γ ⇒ P -> Γ ⇒ P ∨ Q 18 | | SC_or_intror {Γ P Q} : Γ ⇒ Q -> Γ ⇒ P ∨ Q 19 | | SC_or_elim {Γ P Q R} : In (P ∨ Q) Γ -> P :: Γ ⇒ R -> Q :: Γ ⇒ R -> 20 | Γ ⇒ R 21 | | SC_impl_intro {Γ P Q} : P :: Γ ⇒ Q -> Γ ⇒ P ⊃ Q 22 | | SC_impl_elim {Γ P Q R} : In (P ⊃ Q) Γ -> Γ ⇒ P -> Q :: Γ ⇒ R -> 23 | Γ ⇒ R 24 | where "Γ ⇒ P" := (SC_proves Γ P). 25 | 26 | Proposition SC_consistent : ~ (nil ⇒ ⊥). 27 | Proof. 28 | intro. inversion H; repeat match goal with 29 | | H0 : In _ nil |- _ => destruct H0 30 | end. 31 | Qed. 32 | 33 | Global Instance SC_context_extension : 34 | Proper (subcontext ++> eq ==> Basics.impl) SC_proves. 35 | Proof. 36 | intros Γ₁ Γ₂ ? P Q [] ?. clear Q; revert Γ₂ H. 37 | induction H0; eauto using SC_proves. 38 | + intros. apply @SC_and_elim with (P := P) (Q := Q); auto. 39 | apply IHSC_proves. repeat f_equiv. assumption. 40 | + intros. apply @SC_or_elim with (P := P) (Q := Q); auto. 41 | - apply IHSC_proves1. f_equiv; assumption. 42 | - apply IHSC_proves2. f_equiv; assumption. 43 | + intros. apply SC_impl_intro. apply IHSC_proves. f_equiv; assumption. 44 | + intros. apply @SC_impl_elim with (P := P) (Q := Q); auto. 45 | apply IHSC_proves2. f_equiv; assumption. 46 | Qed. 47 | 48 | Inductive subterm : prop atom -> prop atom -> Prop := 49 | | subterm_and_l {P Q} : subterm P (P ∧ Q) 50 | | subterm_and_r {P Q} : subterm Q (P ∧ Q) 51 | | subterm_or_l {P Q} : subterm P (P ∨ Q) 52 | | subterm_or_r {P Q} : subterm Q (P ∨ Q) 53 | | subterm_impl_l {P Q} : subterm P (P ⊃ Q) 54 | | subterm_impl_r {P Q} : subterm Q (P ⊃ Q). 55 | 56 | Lemma subterm_ind_scheme : forall (A : prop atom -> Prop), 57 | (forall P (IHP : forall P', subterm P' P -> A P'), A P) -> 58 | forall P, A P. 59 | Proof. 60 | intros A F; induction P; apply F; inversion 1; auto. 61 | Qed. 62 | 63 | Lemma SC_admits_cut_ext : forall P Q Γ Γ', 64 | Γ ⇒ P -> Γ' ⊆ P :: Γ -> Γ' ⇒ Q -> Γ ⇒ Q. 65 | Proof. 66 | induction P using subterm_ind_scheme. intros. 67 | revert Γ H H0; induction H1; intros. 68 | all: try match goal with 69 | | H : In _ ?Γ, H1 : ?Γ ⊆ _ |- _ => pose proof (H1 _ H) 70 | end. 71 | all: repeat match goal with 72 | | H : In _ (_ :: _) |- _ => destruct H 73 | end. 74 | all: subst. 75 | + assumption. 76 | + apply SC_init; assumption. 77 | + remember ⊥ as R. induction H0; try discriminate; subst. 78 | - apply SC_bot_elim; assumption. 79 | - apply SC_bot_elim; assumption. 80 | - apply (SC_and_elim H0). apply IHSC_proves; trivial. 81 | rewrite H1; prove_subcontext. 82 | - apply (SC_or_elim H0). 83 | * apply IHSC_proves1; trivial. rewrite H1; prove_subcontext. 84 | * apply IHSC_proves2; trivial. rewrite H1; prove_subcontext. 85 | - apply (SC_impl_elim H0); trivial. apply IHSC_proves2; trivial. 86 | rewrite H1; prove_subcontext. 87 | + apply SC_bot_elim; assumption. 88 | + apply SC_top_intro. 89 | + apply SC_and_intro. 90 | - apply IHSC_proves1; assumption. 91 | - apply IHSC_proves2; assumption. 92 | + remember (P0 ∧ Q) as S. pose proof H0; induction H0; try discriminate; subst. 93 | - apply (SC_and_elim H0). apply IHSC_proves; trivial. 94 | * do 2 rewrite <- subcontext_cons_r. assumption. 95 | * rewrite H2. prove_subcontext. 96 | - apply SC_bot_elim. assumption. 97 | - injection HeqS; intros; subst. 98 | eapply (IHP Q); trivial; [ constructor | reflexivity | idtac ]. 99 | eapply (IHP P0); [ constructor | | reflexivity | ]. 100 | * rewrite <- subcontext_cons_r; assumption. 101 | * apply IHSC_proves. 102 | { do 2 rewrite <- subcontext_cons_r. assumption. } 103 | { rewrite H2. prove_subcontext. } 104 | - apply (SC_and_elim H0). apply IHSC_proves0; trivial. 105 | rewrite H2. prove_subcontext. 106 | - apply (SC_or_elim H0). 107 | * apply IHSC_proves1; trivial. rewrite H2. prove_subcontext. 108 | * apply IHSC_proves2; trivial. rewrite H2. prove_subcontext. 109 | - apply (SC_impl_elim H0). 110 | * assumption. 111 | * apply IHSC_proves2; trivial. rewrite H2. prove_subcontext. 112 | + apply (SC_and_elim H3). apply IHSC_proves. 113 | - do 2 rewrite <- subcontext_cons_r. assumption. 114 | - rewrite H2. prove_subcontext. 115 | + apply SC_or_introl. apply IHSC_proves; trivial. 116 | + apply SC_or_intror. apply IHSC_proves; trivial. 117 | + remember (P0 ∨ Q) as S. pose proof H0; induction H0; try discriminate; subst. 118 | - apply (SC_or_elim H0). 119 | * apply IHSC_proves1; trivial. 120 | { rewrite <- subcontext_cons_r. assumption. } 121 | { rewrite H1. prove_subcontext. } 122 | * apply IHSC_proves2; trivial. 123 | { rewrite <- subcontext_cons_r. assumption. } 124 | { rewrite H1. prove_subcontext. } 125 | - apply SC_bot_elim. assumption. 126 | - apply (SC_and_elim H0). apply IHSC_proves; trivial. 127 | rewrite H1. prove_subcontext. 128 | - injection HeqS; intros; subst. 129 | eapply (IHP P0); trivial; [ constructor | reflexivity | ]. 130 | apply IHSC_proves1; trivial. 131 | * rewrite <- subcontext_cons_r. assumption. 132 | * rewrite H1. prove_subcontext. 133 | - injection HeqS; intros; subst. 134 | eapply (IHP Q); trivial; [ constructor | reflexivity | ]. 135 | apply IHSC_proves2; trivial. 136 | * rewrite <- subcontext_cons_r. assumption. 137 | * rewrite H1. prove_subcontext. 138 | - apply (SC_or_elim H0). 139 | * apply IHSC_proves3; trivial. rewrite H1. prove_subcontext. 140 | * apply IHSC_proves4; trivial. rewrite H1. prove_subcontext. 141 | - apply (SC_impl_elim H0); trivial. 142 | apply IHSC_proves4; trivial. rewrite H1. prove_subcontext. 143 | + apply (SC_or_elim H2). 144 | - apply IHSC_proves1. 145 | * rewrite <- subcontext_cons_r. assumption. 146 | * rewrite H1. prove_subcontext. 147 | - apply IHSC_proves2. 148 | * rewrite <- subcontext_cons_r. assumption. 149 | * rewrite H1. prove_subcontext. 150 | + apply SC_impl_intro. apply IHSC_proves; trivial. 151 | - rewrite <- subcontext_cons_r. assumption. 152 | - rewrite H0. prove_subcontext. 153 | + remember (P0 ⊃ Q) as S. pose proof H0; induction H0; try discriminate; subst. 154 | - apply (SC_impl_elim H0). 155 | * apply IHSC_proves1; trivial. 156 | * apply IHSC_proves2. 157 | { rewrite <- subcontext_cons_r. assumption. } 158 | { rewrite H1. prove_subcontext. } 159 | - apply SC_bot_elim; assumption. 160 | - apply (SC_and_elim H0). apply IHSC_proves; trivial. 161 | rewrite H1. prove_subcontext. 162 | - apply (SC_or_elim H0). 163 | * apply IHSC_proves3; trivial. rewrite H1. prove_subcontext. 164 | * apply IHSC_proves4; trivial. rewrite H1. prove_subcontext. 165 | - injection HeqS; intros; subst. clear IHSC_proves. 166 | eapply (IHP Q); [ constructor | | reflexivity | ]. 167 | * eapply (IHP P0); [ constructor | | reflexivity | assumption ]. 168 | apply IHSC_proves1; trivial. 169 | * apply IHSC_proves2; trivial. 170 | { rewrite <- subcontext_cons_r; assumption. } 171 | { rewrite H1. prove_subcontext. } 172 | - apply (SC_impl_elim H0); trivial. apply IHSC_proves4; trivial. 173 | rewrite H1. prove_subcontext. 174 | + apply (SC_impl_elim H2). 175 | - apply IHSC_proves1; trivial. 176 | - apply IHSC_proves2. 177 | * rewrite <- subcontext_cons_r. assumption. 178 | * rewrite H1. prove_subcontext. 179 | Qed. 180 | 181 | Theorem SC_admits_cut : forall Γ P Q, 182 | Γ ⇒ P -> P :: Γ ⇒ Q -> Γ ⇒ Q. 183 | Proof. 184 | intros. eapply SC_admits_cut_ext; (eassumption || reflexivity). 185 | Qed. 186 | 187 | End SequentCalculus. 188 | 189 | Infix "⇒" := SC_proves (no associativity, at level 61). 190 | -------------------------------------------------------------------------------- /Subcontext.v: -------------------------------------------------------------------------------- 1 | Require Export PropLang. 2 | Require Export RelationClasses. 3 | Require Export Morphisms. 4 | Require Export List. 5 | 6 | Section subcontext. 7 | 8 | Context {atom : Type}. 9 | 10 | Definition subcontext (Γ₁ Γ₂ : list (prop atom)) : Prop := 11 | forall P, In P Γ₁ -> In P Γ₂. 12 | Infix "⊆" := subcontext (no associativity, at level 70). 13 | 14 | Global Instance subcontext_preord : PreOrder subcontext. 15 | Proof. 16 | constructor. 17 | + intro Γ. red. trivial. 18 | + intros Γ₁ Γ₂ Γ₃; unfold subcontext. auto. 19 | Qed. 20 | 21 | Lemma subcontext_cons : forall P Γ₁ Γ₂, P :: Γ₁ ⊆ Γ₂ <-> 22 | In P Γ₂ /\ Γ₁ ⊆ Γ₂. 23 | Proof. 24 | split; intros; repeat split. 25 | + apply H; left; reflexivity. 26 | + intros x ?; apply H; right; assumption. 27 | + destruct H. intro x; destruct 1; subst; auto. 28 | Qed. 29 | 30 | Lemma subcontext_cons_r : forall P Γ, Γ ⊆ P :: Γ. 31 | Proof. 32 | intros P Γ x ?; right; assumption. 33 | Qed. 34 | 35 | Global Instance subcontext_cons_proper : 36 | Proper (eq ==> subcontext ++> subcontext) (@cons (prop atom)). 37 | Proof. 38 | intros P Q [] Γ₁ Γ₂ ?. rewrite subcontext_cons; split. 39 | + left; reflexivity. 40 | + rewrite H; apply subcontext_cons_r. 41 | Qed. 42 | 43 | End subcontext. 44 | 45 | Infix "⊆" := subcontext (no associativity, at level 70). 46 | 47 | Ltac prove_In := 48 | match goal with 49 | | |- In ?P (?P :: ?Γ) => left; reflexivity 50 | | |- In ?P (?Q :: ?Γ) => right; prove_In 51 | end. 52 | Ltac prove_subcontext := 53 | match goal with 54 | | |- ?P :: ?Γ ⊆ ?Γ' => rewrite subcontext_cons; split; 55 | [ prove_In | prove_subcontext ] 56 | | |- ?Γ ⊆ ?Γ => reflexivity 57 | | |- ?Γ ⊆ ?P :: ?Γ' => rewrite <- (subcontext_cons_r P Γ'); 58 | prove_subcontext 59 | end. 60 | --------------------------------------------------------------------------------