diff --git a/boolean_completeness/Makefile b/boolean_completeness/Makefile
new file mode 100644
index 0000000..77b94df
--- /dev/null
+++ b/boolean_completeness/Makefile
@@ -0,0 +1,26 @@
+
+COQC=coqc
+COQDOC=coqdoc
+DVIPDF=dvipdf
+
+all: main doc
+ coqc lists.v filters.v classcomp.v
+
+main: pairing
+ coqc lists.v filters.v classcomp.v
+
+pairing: force_look
+ cd pairing; coqc extEqualNat.v misc.v primRec.v cPair.v
+
+doc:
+ $(COQDOC) -d html -g --utf8 --toc --no-index *.v pairing/*.v
+
+doc-pdf:
+ $(COQDOC) -p "\usepackage{hyperref}\hypersetup{colorlinks=true,linkcolor=red}" -o boolean_completeness.dvi --dvi -toc -g *.v pairing/*.v
+ $(DVIPDF) boolean_completeness.dvi
+
+clean:
+ rm -f *.vo *.glob pairing/*.vo pairing/*.glob html/pairing*.html html/classcomp.html html/filters.html html/lists.html html/toc.html html/coqdoc.css *.dvi *.pdf
+
+force_look:
+ true
diff --git a/boolean_completeness/classcomp.v b/boolean_completeness/classcomp.v
new file mode 100644
index 0000000..b76c034
--- /dev/null
+++ b/boolean_completeness/classcomp.v
@@ -0,0 +1,3641 @@
+(**
+ Krivine's constructive proof of completeness for classical
+ first-order logic, following the paper of Berardi and Valentini.
+
+ Uses Russell O'Connor's implementation of the Cantor Pairing
+ function in Coq.
+
+ A part involving list sorting is not fully formalised.
+
+ Danko Ilik, October 2008
+*)
+
+Require Export List.
+Require Import Setoid.
+Require Import Bool.
+Require Import EqNat.
+Require Export Peano_dec.
+Require Import Compare_dec.
+Require Import Max.
+Require Import Le.
+
+(** This imports the proof of the constructive Ultra-filter Theorem *)
+Require Import filters.
+(** This imports the unfinished list-sorting library*)
+Require Import lists.
+
+Set Implicit Arguments.
+
+(** printing ==> $\Rightarrow$ #⇒# *)
+
+(** To build the syntax of formulas, we start with decidable countable sets of constant, function, and predicate symbols. *)
+Module Type classical_completeness_signature.
+ Parameters function predicate constant0 : Set.
+ Parameter function_dec : forall f1 f2:function, {f1 = f2} + {f1 <> f2}.
+ Parameter predicate_dec : forall f1 f2:predicate, {f1 = f2} + {f1 <> f2}.
+ Parameter constant0_dec : forall f1 f2:constant0, {f1 = f2} + {f1 <> f2}.
+ Parameter enum_function : function -> nat.
+ Parameter enum_predicate : predicate -> nat.
+ Parameter enum_constant0 : constant0 -> nat.
+ Parameter enum_function_inj :
+ forall p q, enum_function p = enum_function q -> p = q.
+ Parameter enum_predicate_inj :
+ forall p q, enum_predicate p = enum_predicate q -> p = q.
+ Parameter enum_constant0_inj :
+ forall p q, enum_constant0 p = enum_constant0 q -> p = q.
+End classical_completeness_signature.
+
+Module classical_completeness (ccsig:classical_completeness_signature).
+ Export ccsig.
+
+(**
+ A formula is then defined using the above. There is a special [added] constructor for constants, these are the Henkin constants. There are separate constructors for variables bound by quantifiers [bvar], and free variables [fvar].
+*)
+Inductive formula : Set :=
+| bot : formula
+| imp : formula -> formula -> formula
+| all : formula -> formula
+| atom : predicate -> term -> formula
+with term : Set :=
+| bvar : nat -> term
+| fvar : nat -> term
+| cnst : constant -> term
+| func : function -> term -> term
+with constant : Set :=
+| original : constant0 -> constant
+| added : formula -> constant.
+
+(** 'Opening up' quantifiers, i.e. replacing a de Bruijn variable bound
+ by a quantifier, by a formula. *)
+Fixpoint
+ open_rec (k : nat) (u : term) (t : formula) {struct t} : formula :=
+ match t with
+ | bot => bot
+ | imp t1 t2 => imp (open_rec k u t1) (open_rec k u t2)
+ | all t1 => all (open_rec (S k) u t1)
+ | atom p t1 => atom p (open_rec_term k u t1)
+ end
+with
+ open_rec_term (k : nat) (u : term) (t : term) {struct t} : term :=
+ match t with
+ | bvar i => if beq_nat k i then u else (bvar i)
+ | fvar x => fvar x
+ | cnst c => cnst c
+ | func f t1 => func f (open_rec_term k u t1)
+ end.
+
+(** Substituting the first variable in the term u, by the term t. *)
+Definition open t u := open_rec 0 u t.
+Notation "t ^^ u" := (open t u) (at level 67).
+Notation "t ^ x" := (open t (fvar x)).
+
+(** A formula is [locl] (locally closed) when all [bvar]-s are bound by quantifiers, but there might well be [fvar]-s around. *)
+Definition locl (f:formula) := forall n t, (open_rec n t f) = f.
+
+(** A term is locally-closed if it simply does not have bound
+ variables, but let us define it symmetrically to locl. *)
+Definition loclt (t:term) := forall n t', (open_rec_term n t' t) = t.
+
+Definition locll (Gamma:list formula) := forall B, In B Gamma -> locl B.
+
+Definition notin (x:nat) (L:list nat) := not (In x L).
+Notation "x \notin L" := (notin x L) (at level 69).
+
+(**
+ Natural deduction system for classical predicate logic with cofinite quantification
+*)
+Inductive proof : list formula -> formula -> Prop :=
+| bot_elim : forall Gamma,
+ proof Gamma bot -> forall A, proof Gamma A
+| imp_elim : forall Gamma A B,
+ proof Gamma (imp A B) -> proof Gamma A -> proof Gamma B
+| imp_intro : forall Gamma A B,
+ proof (A::Gamma) B -> proof Gamma (imp A B)
+| dneg : forall Gamma A,
+ proof Gamma (imp (imp A bot) bot) -> proof Gamma A
+| hypo : forall Gamma A,
+ In A Gamma -> proof Gamma A
+| all_elim : forall Gamma A,
+ proof Gamma (all A) -> forall t:term, proof Gamma (A^^t)
+| all_intro : forall Gamma A L, locl (all A) ->
+ (forall x, x \notin L -> proof Gamma (A^x)) ->
+ proof Gamma (all A)
+| weaken : forall Gamma A B, proof Gamma A -> proof (B::Gamma) A.
+
+About all_intro.
+
+Notation "A ==> B" := (imp A B) (at level 55, right associativity).
+Notation "'neg' A" := (imp A bot) (at level 53, right associativity).
+
+(** A general set of formulas *)
+Definition formula_set := formula -> Prop.
+
+(** Definition of a "minimal model", one without standard
+ interpretation of absurdity *)
+Record model (M:formula_set) : Prop := {
+
+ model_dneg : forall A, M (neg neg A ==> A);
+
+ model_imp_faithful1 : forall A B,
+ M (A ==> B) -> (M A -> M B);
+
+ model_imp_faithful2 : forall A B,
+ (M A -> M B) -> M (A ==> B);
+
+ model_all_faithful1 : forall A,
+ M (all A) ->
+ forall t:term, M (A^^t);
+
+ model_all_faithful2 : forall A, locl (all A) ->
+ (forall t:term, loclt t -> M (A^^t)) ->
+ M (all A)
+}.
+
+(** The definition of a "classical" as opposed to a "minimal" model is
+ given, but not used. *)
+Definition model_bot_faithful (M:formula_set) := not (M bot).
+Definition classical_model (M:formula_set) : Prop :=
+ model M /\ model_bot_faithful M.
+
+(** A set of formulas interprets a sequent Gamma|-A if the inclusion of Gamma
+ implies the inclusion of A *)
+Definition interprets (M:formula_set)(Gamma:list formula)(A:formula) :=
+ (forall f, In f Gamma -> M f) -> M A.
+
+(** A sequent is valid if it is true in all models *)
+Definition valid (Gamma:list formula)(A:formula) :=
+ forall M, model M -> interprets M Gamma A.
+
+Section natural_deduction_lemmas.
+ Lemma peirce : forall Gamma P Q, proof Gamma (((P ==> Q) ==> P) ==> P).
+ Proof with (simpl;auto).
+ intros.
+ apply imp_intro...
+ apply dneg...
+ apply imp_intro...
+ apply imp_elim with P...
+ apply hypo...
+ apply imp_elim with (P ==> Q)...
+ apply hypo...
+ apply imp_intro...
+ apply bot_elim...
+ apply imp_elim with P...
+ apply hypo...
+ apply hypo...
+ Defined.
+
+ Lemma proof_imp_trans : forall Gamma x y z,
+ proof Gamma (x ==> y) -> proof Gamma (y ==> z) -> proof Gamma (x ==> z).
+ Proof.
+ intros Gamma x y z Hxy Hyz.
+ apply imp_intro.
+ apply imp_elim with y.
+ apply weaken.
+ assumption.
+ apply imp_elim with x.
+ apply weaken.
+ assumption.
+ apply hypo; simpl; auto.
+ Defined.
+
+ Lemma proof_imp_contrapos : forall Gamma x y,
+ proof Gamma (x ==> y) -> proof Gamma (neg y ==> neg x).
+ Proof.
+ intros.
+ apply imp_intro.
+ apply imp_intro.
+ apply imp_elim with y.
+ apply weaken.
+ apply hypo; simpl; auto.
+ apply imp_elim with x.
+ apply weaken.
+ apply weaken.
+ assumption.
+ apply hypo; simpl; auto.
+ Defined.
+End natural_deduction_lemmas.
+
+(** Some tactics used for building the Lindenbaum algebra below *)
+Ltac impi := apply imp_intro.
+Ltac impe := fun x => apply imp_elim with x.
+Ltac dneg := apply dneg.
+Ltac hypo := apply hypo; simpl; auto.
+Ltac weak := apply weaken.
+Ltac tran := fun x => apply proof_imp_trans with x.
+Ltac contra := apply proof_imp_contrapos.
+
+Lemma formula_dec : forall x y:formula, {x = y}+{x <> y}.
+Proof.
+ fix 1.
+ decide equality.
+ decide equality.
+ decide equality.
+ decide equality.
+ decide equality.
+ apply constant0_dec.
+ apply function_dec.
+ apply predicate_dec.
+Defined.
+
+Lemma constant_dec : forall f1 f2:constant, {f1 = f2} + {f1 <> f2}.
+Proof.
+ fix 1.
+ decide equality.
+ apply constant0_dec.
+ apply formula_dec.
+Defined.
+
+Module Export CBAproof <: CBA.
+
+(** The Lindenbaum Boolean algebra which will be used in the model
+ existence lemma to build a maximal consistent extension of a set of
+ formulas. (the "Universal model") *)
+Section boolean_algebra.
+
+ Let B := formula.
+
+ Definition compl : B -> B.
+ intro x.
+ exact (neg x).
+ Defined.
+
+ Definition meet : B -> B -> B.
+ intros x y.
+ exact (neg (x ==> (neg y))).
+ Defined.
+
+ Definition join : B -> B -> B.
+ intros x y.
+ exact ((neg x) ==> y).
+ Defined.
+
+ Definition top : B.
+ exact (bot ==> bot).
+ Defined.
+
+ (* equality has to be in prop for practical reasons - defining a
+ Coq-setoid *)
+ Definition eq_B (x y:B): Prop :=
+ (exists p:proof nil (x ==> y), True)
+ /\ (exists p:proof nil (y ==> x), True).
+
+ Theorem eq_B_refl : reflexive B eq_B.
+ Proof.
+ red.
+ unfold eq_B.
+ intros.
+ assert (proof nil (x ==> x)).
+ apply imp_intro.
+ apply hypo.
+ simpl.
+ left.
+ reflexivity.
+ firstorder.
+ Defined.
+
+ Theorem eq_B_symm : symmetric B eq_B.
+ Proof.
+ red.
+ unfold eq_B.
+ intros.
+ firstorder.
+ Defined.
+
+ Theorem eq_B_trans : transitive B eq_B.
+ Proof.
+ red.
+ unfold eq_B.
+ intros.
+ set (X:=x) in *.
+ set (Y:=y) in *.
+ set (Z:=z) in *.
+ destruct H as [[pXY _] [pYX _]].
+ destruct H0 as [[pYZ _] [pZY _]].
+ assert (pXZ:=proof_imp_trans pXY pYZ).
+ assert (pZX:=proof_imp_trans pZY pYX).
+ firstorder.
+ Defined.
+
+ Notation "x == y" := (eq_B x y) (at level 70, no associativity).
+
+ Add Relation B eq_B
+ reflexivity proved by eq_B_refl
+ symmetry proved by eq_B_symm
+ transitivity proved by eq_B_trans
+ as eq_B_relation.
+
+ Add Morphism join with signature eq_B ==> eq_B ==> eq_B as join_morphism.
+ Proof.
+ unfold eq_B; try red.
+ intros x y H x0 y0 H0.
+ generalize dependent (x0).
+ generalize dependent (x).
+ generalize dependent (y).
+ generalize dependent (y0).
+ clear.
+ intros W X Y H1 Z H2.
+ destruct H1 as [[pYX _] [pXY _]].
+ destruct H2 as [[pZW _] [pWZ _]].
+ split.
+ assert (proof nil ((neg Y ==> Z) ==> (neg X ==> W))).
+ apply proof_imp_trans with (neg Y ==> W).
+ apply imp_intro.
+ apply imp_intro.
+ apply imp_elim with Z.
+ apply weaken.
+ apply weaken.
+ assumption.
+ apply imp_elim with (neg Y).
+ apply weaken.
+ apply hypo.
+ simpl; auto.
+ apply hypo.
+ simpl; auto.
+ apply imp_intro.
+ apply imp_intro.
+ apply imp_elim with (neg Y).
+ apply weaken.
+ apply hypo.
+ simpl; auto.
+ apply imp_elim with (neg X).
+ apply weaken.
+ apply weaken.
+ apply proof_imp_contrapos.
+ assumption.
+ apply hypo.
+ simpl; auto.
+ exists H.
+ auto.
+ (* completelly the same as the above proof *)
+ assert (proof nil ((neg X ==> W) ==> (neg Y ==> Z))).
+ apply proof_imp_trans with (neg X ==> Z).
+ apply imp_intro.
+ apply imp_intro.
+ apply imp_elim with W.
+ apply weaken.
+ apply weaken.
+ assumption.
+ apply imp_elim with (neg X).
+ apply weaken.
+ apply hypo.
+ simpl; auto.
+ apply hypo.
+ simpl; auto.
+ apply imp_intro.
+ apply imp_intro.
+ apply imp_elim with (neg X).
+ apply weaken.
+ apply hypo.
+ simpl; auto.
+ apply imp_elim with (neg Y).
+ apply weaken.
+ apply weaken.
+ apply proof_imp_contrapos.
+ assumption.
+ apply hypo.
+ simpl; auto.
+ exists H.
+ auto.
+ Defined.
+
+ Add Morphism meet with signature eq_B ==> eq_B ==> eq_B as meet_morphism.
+ Proof.
+ unfold eq_B.
+ intros x y H x0 y0 H0.
+ generalize dependent (x0).
+ generalize dependent (x).
+ generalize dependent (y).
+ generalize dependent (y0).
+ clear.
+ intros W X Y H1 Z H2.
+ destruct H1 as [[pYX _] [pXY _]].
+ destruct H2 as [[pZW _] [pWZ _]].
+ split.
+ assert (proof nil (neg (Y ==> neg Z) ==> neg (X ==> neg W))).
+ impi.
+ impi.
+ impe (Y ==> neg Z).
+ weak;hypo.
+ impi.
+ impe (neg W).
+ apply proof_imp_contrapos.
+ weak;weak;weak;assumption.
+ impe X.
+ hypo;weak.
+ impe Y.
+ weak;weak;weak;assumption.
+ hypo.
+ exists H; auto.
+ assert (proof nil (neg (X ==> neg W) ==> neg (Y ==> neg Z))).
+ apply proof_imp_contrapos.
+ impi.
+ tran (neg Z).
+ tran Y.
+ weak;assumption.
+ hypo.
+ contra.
+ weak;assumption.
+ exists H; auto.
+ Defined.
+
+ Lemma id_B_dec : forall x y : B, {x = y}+{x <> y}.
+ intros.
+ apply formula_dec.
+ Defined.
+
+ Lemma id_B_dec_right : forall (x y:B), x<>y ->
+ exists H:x<>y, id_B_dec x y = right (x=y) H.
+ Proof.
+ intros.
+ unfold id_B_dec.
+ case (formula_dec x y).
+ (* case 1 *)
+ intros eqxy.
+ congruence.
+ (* case 2 *)
+ intro.
+ exists n.
+ reflexivity.
+ Defined.
+
+ Lemma id_B_dec_left : forall x:B, exists H:x=x,
+ id_B_dec x x = left (x<>x) H.
+ Proof.
+ intros.
+ unfold id_B_dec.
+ case (formula_dec x x).
+ (* case 1 *)
+ intro.
+ exists e.
+ reflexivity.
+ (* case 2 *)
+ intros neqxy.
+ congruence.
+ Defined.
+
+ Lemma join_idem : forall x, join x x == x.
+ intro x.
+ unfold eq_B.
+ generalize dependent (x).
+ clear.
+ intro f.
+ split.
+ assert (proof nil ((neg f ==> f) ==> f)).
+ impi.
+ dneg.
+ impi.
+ impe f.
+ hypo.
+ impe (neg f).
+ weak.
+ hypo.
+ hypo.
+ firstorder.
+ assert (proof nil (f ==> (neg f ==> f))).
+ impi.
+ impi.
+ weak.
+ hypo.
+ firstorder.
+ Defined.
+
+ Lemma meet_idem : forall x, meet x x == x.
+ Proof.
+ intro.
+ unfold eq_B.
+ simpl proj1_sig.
+ generalize dependent (x).
+ intro f.
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil (neg (f ==> neg f) ==> f)).
+ impi.
+ dneg.
+ impi.
+ impe (f==>neg f).
+ weak; hypo.
+ impi.
+ weak; hypo.
+ exists H; auto.
+ assert (proof nil (f ==> neg (f ==> neg f))).
+ impi.
+ impi.
+ impe f.
+ impe f.
+ hypo.
+ weak; hypo.
+ weak; hypo.
+ exists H; auto.
+ Defined.
+
+ Lemma join_comm : forall x y, join x y == join y x.
+ Proof.
+ intros.
+ unfold eq_B.
+ generalize dependent (x).
+ generalize dependent (y).
+ intros X Y.
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil ((neg Y ==> X) ==> (neg X ==> Y))).
+ impi.
+ impi.
+ dneg.
+ impi.
+ impe X.
+ weak; hypo.
+ impe (neg Y).
+ weak; weak; hypo.
+ hypo.
+ exists H; auto.
+ assert (proof nil ((neg X ==> Y) ==> (neg Y ==> X))).
+ impi.
+ impi.
+ dneg.
+ impi.
+ impe Y.
+ weak; hypo.
+ impe (neg X).
+ weak; weak; hypo.
+ hypo.
+ exists H; auto.
+ Defined.
+
+ Lemma meet_comm : forall x y, meet x y == meet y x.
+ Proof.
+ intros.
+ unfold eq_B.
+ generalize dependent (x).
+ generalize dependent (y).
+ intros X Y.
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil (neg (Y ==> neg X) ==> neg (X ==> neg Y))).
+ impi.
+ impi.
+ impe (Y ==> neg X).
+ weak; hypo.
+ impi.
+ impi.
+ impe Y.
+ impe X.
+ weak; weak; hypo.
+ hypo.
+ weak; hypo.
+ exists H; auto.
+ assert (proof nil (neg (X ==> neg Y) ==> neg (Y ==> neg X))).
+ impi; impi.
+ impe (X ==> neg Y).
+ hypo; weak.
+ impi;impi.
+ impe X.
+ impe Y.
+ weak; weak; hypo.
+ hypo.
+ weak; hypo.
+ exists H; auto.
+ Defined.
+
+ Lemma join_assoc : forall x y z, join x (join y z) == join (join x y) z.
+ Proof.
+ intros.
+ unfold eq_B.
+ generalize dependent (x).
+ generalize dependent (y).
+ generalize dependent (z).
+ intros X Y Z.
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil ((neg Z ==> (neg Y ==> X)) ==> (neg (neg Z ==> Y) ==> X))).
+ impi.
+ impi.
+ impe (neg Y).
+ impe (neg Z).
+ weak; hypo.
+ impi.
+ impe (neg Z ==> Y).
+ weak; hypo.
+ impi.
+ apply bot_elim.
+ impe Z.
+ hypo.
+ weak; hypo.
+ impi.
+ impe (neg Z ==> Y).
+ weak; hypo.
+ impi.
+ weak; hypo.
+ exists H; auto.
+ assert (proof nil ((neg (neg Z ==> Y) ==> X) ==> (neg Z ==> (neg Y ==> X)))).
+ impi;impi;impi.
+ impe (neg (neg Z ==> Y)).
+ weak; weak; hypo.
+ impi.
+ impe Y.
+ weak; hypo.
+ impe (neg Z).
+ hypo.
+ weak;weak;hypo.
+ exists H; auto.
+ Defined.
+
+ Lemma meet_assoc : forall x y z, meet x (meet y z) == meet (meet x y) z.
+ Proof.
+ intros.
+ unfold eq_B.
+ generalize dependent (x).
+ generalize dependent (y).
+ generalize dependent (z).
+ intros X Y Z.
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (
+ proof nil
+ (neg (Z ==> neg neg (Y ==> neg X)) ==>
+ neg (neg (Z ==> neg Y) ==> neg X))
+ ).
+ contra.
+ impi;impi.
+ impi.
+ impe (Y ==> neg X).
+ hypo.
+ impi.
+ impe (neg (Z ==> neg Y)).
+ weak;weak;weak;hypo.
+ impi.
+ impe Y.
+ impe Z.
+ hypo.
+ weak;weak;weak;hypo.
+ weak;hypo.
+ exists H; auto.
+ assert (
+ proof nil
+ (neg (neg (Z ==> neg Y) ==> neg X) ==>
+ neg (Z ==> neg neg (Y ==> neg X)))
+ ).
+ contra.
+ impi.
+ contra.
+ impi.
+ impi.
+ impi.
+ impe (neg (Y ==> neg X)).
+ impe Z.
+ weak;weak;weak;hypo.
+ weak;hypo.
+ impi.
+ impe X.
+ impe Y.
+ hypo.
+ weak;hypo.
+ weak;weak;weak;hypo.
+ exists H; auto.
+ Defined.
+
+ Lemma meet_absorb : forall x y, meet x (join x y) == x.
+ Proof.
+ intros.
+ unfold eq_B.
+ simpl proj1_sig.
+ generalize dependent (x).
+ generalize dependent (y).
+ intros X Y.
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil (neg (Y ==> neg (neg Y ==> X)) ==> Y)).
+ impi.
+ dneg.
+ impi.
+ impe (Y ==> neg (neg Y ==> X)).
+ weak; hypo.
+ impi.
+ apply bot_elim.
+ impe Y.
+ weak; hypo.
+ hypo.
+ exists H; auto.
+ assert (proof nil (Y ==> neg (Y ==> neg (neg Y ==> X)))).
+ impi;impi.
+ impe (neg Y ==> X).
+ impe Y.
+ hypo.
+ weak; hypo.
+ impi.
+ apply bot_elim.
+ impe Y.
+ hypo.
+ weak; weak; hypo.
+ exists H; auto.
+ Defined.
+
+ Lemma join_absorb : forall x y, join x (meet x y) == x.
+ Proof.
+ intros.
+ unfold eq_B.
+ simpl proj1_sig.
+ generalize dependent (x).
+ generalize dependent (y).
+ intros X Y.
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil ((neg Y ==> neg (Y ==> neg X)) ==> Y)).
+ impi.
+ dneg.
+ impi.
+ impe (Y ==> neg X).
+ impe (neg Y).
+ weak; hypo.
+ hypo.
+ impi.
+ apply bot_elim.
+ impe Y.
+ weak; hypo.
+ hypo.
+ exists H; auto.
+ assert (proof nil (Y ==> (neg Y ==> neg (Y ==> neg X)))).
+ impi; impi.
+ apply bot_elim.
+ impe Y.
+ hypo.
+ weak; hypo.
+ exists H; auto.
+ Defined.
+
+ Lemma join_distrib : forall x y z, join (meet x y) z == meet (join x z) (join y z).
+ Proof.
+ intros.
+ unfold eq_B.
+ simpl proj1_sig.
+ generalize dependent (x).
+ generalize dependent (y).
+ generalize dependent (z).
+ intros X Y Z.
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil
+ ((neg neg (Z ==> neg Y) ==> X) ==>
+ neg ((neg Z ==> X) ==> neg (neg Y ==> X)))).
+ impi;impi.
+ impe (neg Y ==> X).
+ impe (neg Z ==> X).
+ hypo.
+ impi.
+ impe (neg neg (Z ==> neg Y)).
+ weak;weak;hypo.
+ impi.
+ impe (Z ==> neg Y).
+ hypo.
+ impi.
+ apply bot_elim.
+ impe Z.
+ weak;weak;hypo.
+ hypo.
+ impi.
+ impe (neg neg (Z ==> neg Y)).
+ weak;weak;hypo.
+ impi.
+ impe (Z ==> neg Y).
+ hypo.
+ impi.
+ weak;weak;hypo.
+ exists H; auto.
+ assert (proof nil
+ (neg ((neg Z ==> X) ==> neg (neg Y ==> X)) ==>
+ (neg neg (Z ==> neg Y) ==> X))).
+ impi.
+ impi.
+ dneg.
+ impi.
+ impe ((neg Z ==> X) ==> neg (neg Y ==> X)).
+ weak;hypo.
+ impi.
+ impi.
+ impe X.
+ weak;weak;hypo.
+ dneg.
+ impi.
+ impe (neg (Z ==> neg Y)).
+ weak;weak;weak;hypo.
+ impi.
+ impe Z.
+ impi.
+ impe (X).
+ weak;weak;hypo.
+ impe (neg Y).
+ weak;weak;weak;hypo.
+ impe Z.
+ weak;hypo.
+ hypo.
+ (* the second subgoal by LEM, Z or not Z*)
+ dneg.
+ impi.
+ impe X.
+ weak;weak;hypo.
+ impe (neg Z).
+ weak;weak;weak;hypo.
+ hypo.
+ exists H; auto.
+ Defined.
+
+ Lemma meet_bott : forall x, meet bot x == bot.
+ Proof.
+ intros.
+ unfold eq_B.
+ generalize dependent (x).
+ intros X .
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil (neg neg (bot ==> neg X))).
+ impi.
+ impe (bot ==> neg X).
+ hypo.
+ impi.
+ impi.
+ weak;hypo.
+ exists H;auto.
+ assert (proof nil (bot ==> neg (bot ==> neg X))).
+ impi.
+ impi.
+ weak;hypo.
+ exists H;auto.
+ Defined.
+
+ Lemma join_bott : forall x, join bot x == x.
+ Proof.
+ intros.
+ unfold eq_B.
+ generalize dependent (x).
+ intros X .
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil ((neg bot ==> X) ==> X)).
+ impi.
+ impe (neg bot).
+ hypo.
+ impi.
+ hypo.
+ exists H;auto.
+ assert (proof nil (X ==> (neg bot ==> X))).
+ impi.
+ impi.
+ weak;hypo.
+ exists H;auto.
+ Defined.
+
+ Lemma meet_top : forall x, meet top x == x.
+ Proof.
+ intros.
+ unfold eq_B.
+ generalize dependent (x).
+ intros X .
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil (neg (neg bot ==> neg X) ==> X)).
+ impi.
+ dneg.
+ impi.
+ impe (neg bot ==> neg X).
+ weak;hypo.
+ impi.
+ weak; hypo.
+ exists H;auto.
+ assert (proof nil (X ==> neg (neg bot ==> neg X))).
+ impi.
+ impi.
+ impe (neg bot).
+ impi.
+ impe X.
+ impe (neg bot).
+ weak;hypo.
+ hypo.
+ weak;weak;hypo.
+ impi.
+ hypo.
+ exists H;auto.
+ Defined.
+
+ Lemma join_top : forall x, join top x == top.
+ Proof.
+ intros.
+ unfold eq_B.
+ generalize dependent (x).
+ intros X .
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil ((neg neg bot ==> X) ==> neg bot)).
+ impi.
+ impi.
+ hypo.
+ exists H;auto.
+ assert (proof nil (neg bot ==> (neg neg bot ==> X))).
+ impi.
+ impi.
+ dneg.
+ impi.
+ impe (neg bot).
+ weak;hypo.
+ weak;weak;hypo.
+ exists H;auto.
+ Defined.
+
+ Lemma meet_compl : forall x, meet x (compl x) == bot.
+ Proof.
+ intros.
+ unfold eq_B.
+ generalize dependent (x).
+ intros X .
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil (neg neg (X ==> neg neg X))).
+ impi.
+ impe (X ==> neg neg X).
+ hypo.
+ impi.
+ impi.
+ impe X.
+ hypo.
+ weak;hypo.
+ exists H;auto.
+ assert (proof nil (bot ==> neg (X ==> neg neg X))).
+ impi.
+ impi.
+ weak; hypo.
+ exists H;auto.
+ Defined.
+
+ Lemma join_compl : forall x, join x (compl x) == top.
+ Proof.
+ intros.
+ unfold eq_B.
+ generalize dependent (x).
+ intros X .
+ clear.
+ split.
+ (* this all should be done by one Ltac command *)
+ assert (proof nil ((neg X ==> neg X) ==> neg bot)).
+ impi.
+ impi.
+ hypo.
+ exists H;auto.
+ assert (proof nil (neg bot ==> (neg X ==> neg X))).
+ impi.
+ impi.
+ hypo.
+ exists H;auto.
+ Defined.
+
+ Lemma meet_distrib : forall x y z, meet (join x y) z == join (meet x z) (meet y z).
+ Proof.
+ (* one distributive law can be derived from the other one - no need
+ to make a natural deduction derivation *)
+ intros b c a.
+ rewrite (meet_comm (join b c) a).
+ rewrite (meet_comm b a).
+ rewrite (meet_comm c a).
+ rewrite join_distrib.
+ rewrite join_absorb.
+ rewrite (join_comm b (meet a c)).
+ rewrite join_distrib.
+ rewrite meet_assoc.
+ rewrite meet_absorb.
+ rewrite join_comm.
+ reflexivity.
+ Defined.
+End boolean_algebra.
+
+(** To use the completion of filters from filters.v, we also need an
+enumeration of the elements of the Boolean algebra, which is achieved
+by borrowing code about the Cantor pairing function from Russell
+O'Connor's formalisation of the incompleteness theorem *)
+Section Enumeration.
+ Add LoadPath "pairing".
+ Require Import cPair.
+
+ Definition enump := fun p => cPair 11 (enum_predicate p).
+ Definition enumc0 := fun c => cPair 12 (enum_constant0 c).
+ Definition enumfunc := fun f => cPair 13 (enum_function f).
+
+ Fixpoint enumf (f:formula) : nat :=
+ match f with
+ | (atom p t) => cPair 1 (cPair (enump p) (enumt t))
+ | (all g) => cPair 2 (enumf g)
+ | (imp g h) => cPair 3 (cPair (enumf g) (enumf h))
+ | bot => 4
+ end
+ with enumt (t:term) : nat :=
+ match t with
+ | (func phi t') => cPair 5 (cPair (enumfunc phi) (enumt t'))
+ | (cnst c) => cPair 6 (enumc c)
+ | (fvar x) => cPair 7 x
+ | (bvar x) => cPair 8 x
+ end
+ with enumc (c:constant) : nat :=
+ match c with
+ | (added x) => cPair 9 (enumf x)
+ | (original x) => cPair 10 (enumc0 x)
+ end.
+
+ (* Eval compute in (enumf (imp bot bot)). *)
+
+ Scheme Induction for formula Sort Prop
+ with Induction for term Sort Prop
+ with Induction for constant Sort Prop.
+
+ Combined Scheme ftc_ind from formula_ind, term_ind, constant_ind.
+
+ Theorem countable_ftc :
+ (forall f g, enumf f = enumf g -> f = g)
+ /\ (forall t s, enumt t = enumt s -> t = s)
+ /\ (forall c k, enumc c = enumc k -> c = k).
+ Proof.
+ apply ftc_ind.
+
+ intros g Hg.
+ destruct g.
+ reflexivity.
+ simpl in Hg.
+ unfold cPair in Hg.
+ simpl in Hg.
+ discriminate.
+ simpl in Hg.
+ unfold cPair in Hg.
+ simpl in Hg.
+ rewrite <- Plus.plus_Snm_nSm in Hg.
+ simpl in Hg.
+ discriminate.
+ simpl in Hg.
+ unfold cPair in Hg.
+ simpl in Hg.
+ discriminate.
+
+ intros f1 Hf1 f2 Hf2 g Hg.
+ destruct g.
+ simpl in Hg.
+ unfold cPair in Hg.
+ simpl in Hg.
+ discriminate.
+ simpl in Hg.
+ apply cPairInj2 in Hg.
+ assert (Hg' := Hg).
+ apply cPairInj1 in Hg.
+ apply cPairInj2 in Hg'.
+ rewrite Hf1 with g1; [ idtac | assumption].
+ rewrite Hf2 with g2; [ idtac | assumption].
+ reflexivity.
+ simpl in Hg.
+ apply cPairInj1 in Hg.
+ congruence.
+ simpl in Hg.
+ apply cPairInj1 in Hg.
+ congruence.
+
+ intros f1 Hf1 g Hg.
+ destruct g.
+ simpl in Hg.
+ unfold cPair in Hg.
+ simpl in Hg.
+ rewrite <- Plus.plus_Snm_nSm in Hg.
+ simpl in Hg.
+ discriminate.
+ simpl in Hg.
+ apply cPairInj1 in Hg.
+ congruence.
+ simpl in Hg.
+ apply cPairInj2 in Hg.
+ rewrite Hf1 with g; [ idtac | assumption].
+ reflexivity.
+ simpl in Hg.
+ apply cPairInj1 in Hg.
+ congruence.
+
+ intros p t Ht g Hg.
+ destruct g.
+ simpl in Hg.
+ unfold cPair in Hg.
+ simpl in Hg.
+ discriminate.
+ simpl in Hg.
+ apply cPairInj1 in Hg.
+ congruence.
+ simpl in Hg.
+ apply cPairInj1 in Hg.
+ congruence.
+ simpl in Hg.
+ apply cPairInj2 in Hg.
+ assert (Hg' := Hg).
+ apply cPairInj1 in Hg.
+ apply cPairInj2 in Hg'.
+ unfold enump in Hg.
+ apply cPairInj2 in Hg.
+ apply enum_predicate_inj in Hg.
+ apply Ht in Hg'.
+ rewrite Hg.
+ rewrite Hg'.
+ reflexivity.
+
+ intros n s Hs.
+ destruct s.
+ simpl in Hs.
+ apply cPairInj2 in Hs.
+ rewrite Hs.
+ reflexivity.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+
+ intros n s Hs.
+ destruct s.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+ simpl in Hs.
+ apply cPairInj2 in Hs.
+ rewrite Hs.
+ reflexivity.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+
+ intros c Hc s Hs.
+ destruct s.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+ simpl in Hs.
+ apply cPairInj2 in Hs.
+ rewrite Hc with c0.
+ reflexivity.
+ assumption.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+
+ intros f t Ht s Hs.
+ destruct s.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+ simpl in Hs.
+ apply cPairInj1 in Hs.
+ congruence.
+ simpl in Hs.
+ apply cPairInj2 in Hs.
+ assert (Hs' := Hs).
+ apply cPairInj1 in Hs.
+ apply cPairInj2 in Hs'.
+ rewrite Ht with s; [ | assumption].
+ unfold enumfunc in Hs.
+ apply cPairInj2 in Hs.
+ apply enum_function_inj in Hs.
+ rewrite Hs.
+ reflexivity.
+
+ intros c k Hk.
+ destruct k.
+ simpl in Hk.
+ apply cPairInj2 in Hk.
+ unfold enumc0 in Hk.
+ apply cPairInj2 in Hk.
+ apply enum_constant0_inj in Hk.
+ rewrite Hk.
+ reflexivity.
+ simpl in Hk.
+ apply cPairInj1 in Hk.
+ congruence.
+
+ intros f Hf k Hk.
+ destruct k.
+ simpl in Hk.
+ apply cPairInj1 in Hk.
+ congruence.
+ simpl in Hk.
+ apply cPairInj2 in Hk.
+ rewrite Hf with f0; [|assumption].
+ reflexivity.
+ Qed.
+
+ Definition enum := enumf.
+
+ Definition countable : forall x y, enum x = enum y -> x = y
+ := proj1 countable_ftc.
+
+End Enumeration.
+
+Definition eq_B_join_morph := join_morphism_Proper.
+Definition eq_B_meet_morph := meet_morphism_Proper.
+Definition bott := bot.
+Definition B := formula.
+
+End CBAproof.
+
+(** Various lemmas that have to do with general lists or their
+ interaction with proofs *)
+Section list_proof_lemmas.
+
+ (* a finite list of formulas is included in a set of formulas *)
+ Definition included (Gamma:list formula)(T:formula_set) :=
+ forall f, In f Gamma -> T f.
+
+ Lemma nil_included : forall Ax, included nil Ax.
+ Proof.
+ red.
+ simpl.
+ intros.
+ absurd (False); auto.
+ Qed.
+
+ Lemma nil_lem1 : forall l:list formula, incl nil l.
+ Proof.
+ red.
+ simpl.
+ intros.
+ absurd (False); auto.
+ Qed.
+
+ Lemma included_lem1 : forall l1 l2:list formula, forall Ax:formula_set,
+ included l1 Ax -> included l2 Ax -> included (l1++l2) Ax.
+ Proof.
+ unfold included.
+ intros.
+ destruct (in_app_or l1 l2 f H1); auto.
+ Qed.
+
+ Lemma weaken_lem1 : forall Gamma Delta A, incl Gamma Delta ->
+ proof Gamma A -> proof Delta A.
+ Proof.
+ intros.
+ generalize dependent Delta.
+ induction H0.
+
+ intros.
+ apply bot_elim;auto.
+
+ intros.
+ apply imp_elim with A; auto.
+
+ intros.
+ apply imp_intro; auto.
+ apply IHproof.
+ clear -H.
+ unfold incl in *.
+ simpl.
+ intros; intuition.
+
+ auto using dneg.
+
+ auto using hypo.
+
+ auto using all_elim.
+
+ eauto using all_intro.
+
+ intros.
+ assert (incl Gamma Delta).
+ clear -H.
+ unfold incl in *.
+ simpl in *.
+ intros; intuition.
+ auto.
+ Qed.
+
+End list_proof_lemmas.
+
+(** A couple of lemmas about locally closed formulas *)
+Section locl_lemmas.
+ Lemma locl_all_neg : forall A, locl (all A) -> locl (all (neg A)).
+ Proof.
+ unfold locl.
+ simpl.
+ intros A H.
+ assert (H' : forall (n : nat) (t : term), (open_rec (S n) t A) = A).
+ intros.
+ assert (H1 := H n t).
+ congruence.
+ intros.
+ rewrite H'.
+ reflexivity.
+ Qed.
+
+ Lemma locl_henkin : forall A,
+ locl (all A) -> locl (all (A ==> all A)).
+ Proof.
+ unfold locl.
+ simpl.
+ intros A H.
+ assert (H' : forall (n : nat) (t : term), (open_rec (S n) t A) = A).
+ intros.
+ assert (H1 := H n t).
+ congruence.
+ intros.
+ rewrite H'.
+ rewrite H'.
+ reflexivity.
+ Qed.
+
+ Lemma locl_henkin_neg : forall A,
+ locl (all A) -> locl (all (neg (A ==> all A))).
+ Proof.
+ unfold locl.
+ simpl.
+ intros A H.
+ assert (H' : forall (n : nat) (t : term), (open_rec (S n) t A) = A).
+ intros.
+ assert (H1 := H n t).
+ congruence.
+ intros.
+ rewrite H'.
+ rewrite H'.
+ reflexivity.
+ Qed.
+End locl_lemmas.
+
+(** This section defines a fixpoint [c_appears] which determines if a
+ given constant appears in the formula and then goes on to prove
+ that [c_appears f (added (all f)) = false], i.e. that a formula
+ cannot contain an added (Henkin) constant indexed by itself. This
+ is obvious, but the proof has to go on the induction of the depth
+ of the formula.
+
+ Another fixpoint [added_cnsts] is also defined, to check if a
+ formula contains _any_ added constants and is connected to
+ [c_appears].
+*)
+Section constants_in_formulas.
+ Fixpoint c_appears (t:formula)(c:constant) {struct t} : bool :=
+ match t with
+ | bot => false
+ | imp t1 t2 => orb (c_appears t1 c) (c_appears t2 c)
+ | all t1 => c_appears t1 c
+ | atom p t1 => c_appears_term t1 c
+ end
+ with
+ c_appears_term (t:term)(c:constant) {struct t} : bool :=
+ match t with
+ | bvar i => false
+ | fvar x => false
+ | cnst k => if (constant_dec k c) then true else false
+ | func f t1 => c_appears_term t1 c
+ end.
+
+ (** c_appears applied to a list *)
+ Fixpoint c_appears_l (l:list formula)(c:constant) {struct l} : bool :=
+ match l with
+ | (cons x x0) => orb (c_appears x c) (c_appears_l x0 c)
+ | nil => false
+ end.
+
+ Fixpoint depth (f:formula) : nat :=
+ match f with
+ | (atom p t) => depth_term t
+ | (all g) => S (depth g)
+ | (imp g h) => S (max (depth g) (depth h))
+ | bot => 1
+ end
+ with
+ depth_term (t:term) : nat :=
+ match t with
+ | (func f t') => S (depth_term t')
+ | (cnst c) =>
+ match c with
+ | original k => 1
+ | added f => S (depth f)
+ end
+ | (fvar x) => 1
+ | (bvar x) => 1
+ end.
+
+ Lemma bb''' : forall f g, depth f <= depth g -> c_appears f (added g) = false.
+ Proof.
+ intros f g H.
+ induction f.
+ simpl.
+ reflexivity.
+ simpl.
+ simpl in H.
+ rewrite max_SS in H.
+ assert (depth f1 <= depth g).
+ apply le_trans with (S (depth f1)).
+ auto.
+ eapply le_trans.
+ apply le_max_l.
+ apply H.
+ assert (depth f2 <= depth g).
+ apply le_trans with (S (depth f2)).
+ auto.
+ eapply le_trans.
+ apply le_max_r.
+ apply H.
+ rewrite IHf1; auto.
+ simpl in H.
+ simpl.
+ rewrite IHf.
+ reflexivity.
+ apply le_trans with (S (depth f)).
+ auto.
+ assumption.
+ induction t.
+ simpl in *.
+ auto.
+ simpl in *.
+ auto.
+ simpl.
+ destruct (constant_dec c (added g)).
+ rewrite e in H.
+ simpl in H.
+ contradict H.
+ auto with arith.
+ reflexivity.
+ simpl.
+ simpl in IHt.
+ rewrite IHt.
+ reflexivity.
+ simpl in H.
+ apply le_trans with (S (depth_term t)).
+ auto.
+ assumption.
+ Qed.
+
+ Theorem c_appears_thm : forall f, c_appears f (added (all f)) = false.
+ Proof.
+ intro f.
+ apply bb'''.
+ simpl.
+ auto.
+ Qed.
+
+ Lemma c_appears_l_app : forall l1 l2 c,
+ c_appears_l l1 c = false /\ c_appears_l l2 c = false ->
+ c_appears_l (l1++l2) c = false.
+ Proof.
+ induction l1;
+ intros l2 c [H1 H2].
+
+ simpl.
+ assumption.
+
+ simpl.
+ rewrite IHl1.
+ simpl in H1.
+ apply orb_false_elim in H1.
+ destruct H1 as [H11 _].
+ rewrite H11.
+ auto.
+ rewrite H2.
+ simpl in H1.
+ apply orb_false_elim in H1.
+ destruct H1 as [_ H12].
+ rewrite H12.
+ auto.
+ Qed.
+
+ Fixpoint added_cnsts (t:term) : bool :=
+ match t with
+ | (func f t') => added_cnsts t'
+ | (cnst c) =>
+ match c with
+ | (added k) => true
+ | (original k) => false
+ end
+ | (fvar x) => false
+ | (bvar x) => false
+ end.
+
+ Fixpoint added_cnsts_f (f:formula) : bool :=
+ match f with
+ | (atom p t) => added_cnsts t
+ | (all g) => added_cnsts_f g
+ | (imp g h) => added_cnsts_f g || added_cnsts_f h
+ | bot => false
+ end.
+
+ Lemma added_cnsts_c_appears : forall t k,
+ added_cnsts t = false -> c_appears_term t (added k) = false.
+ Proof.
+ induction t; simpl in *; auto.
+ intros k H.
+ destruct c.
+ simpl.
+ auto.
+ congruence.
+ Qed.
+
+ (* the formula counterpart to added_cnsts_c_appears (which works
+ only on terms) *)
+ Lemma added_cnsts_c_appears' : forall f g,
+ added_cnsts_f f = false -> c_appears f (added g) = false.
+ Proof.
+ induction f; simpl.
+
+ auto.
+
+ intros.
+ apply orb_false_elim in H.
+ destruct H.
+ rewrite IHf1; auto.
+ rewrite IHf2; auto.
+
+ auto.
+
+ apply added_cnsts_c_appears.
+ Qed.
+End constants_in_formulas.
+
+(** This section provides that if there is a derivation Gamma |- A,
+ then there is a derivation Gamma{x/c} |- A{x/c}, where {x/c} is a
+ substitution of the constant c by the free variable x. This lemma
+ is needed in [henkin_equiconsistent].
+*)
+Section vanDalen_thm283.
+ (* Substituting a constant by a term *)
+ Fixpoint
+ c2t (t:formula)(c:constant)(x:term) {struct t} : formula :=
+ match t with
+ | bot => bot
+ | imp t1 t2 => imp (c2t t1 c x) (c2t t2 c x)
+ | all t1 => all (c2t t1 c x)
+ | atom p t1 => atom p (c2t_term t1 c x)
+ end
+ with
+ c2t_term (t:term)(c:constant)(x:term) {struct t} : term :=
+ match t with
+ | bvar i => (bvar i)
+ | fvar x => fvar x
+ | cnst k => if (constant_dec k c) then x else (cnst k)
+ | func f t1 => func f (c2t_term t1 c x)
+ end.
+
+ (* c2t applied to a list *)
+ Fixpoint c2tl (l:list formula)(c:constant)(v:term) {struct l} :=
+ match l with
+ | (cons x x0) => (c2t x c v) :: (c2tl x0 c v)
+ | nil => nil
+ end.
+
+ Lemma c2t_idem : forall (f:formula)(c:constant)(x:term),
+ c_appears f c = false -> c2t f c x = f.
+ Proof.
+ induction f.
+ simpl.
+ auto.
+ simpl.
+ intros.
+ apply orb_false_elim in H.
+ destruct H as [H1 H2].
+ rewrite IHf1; auto.
+ rewrite IHf2; auto.
+ simpl.
+ intros.
+ rewrite IHf; auto.
+ induction t.
+ simpl.
+ auto.
+ simpl.
+ auto.
+ simpl.
+ intros.
+ destruct (constant_dec c c0).
+ congruence.
+ auto.
+ simpl in *.
+ intros.
+ assert (c2t_term t c x = t).
+ assert (IH := IHt c x H).
+ clear IHt.
+ congruence.
+ rewrite H0.
+ reflexivity.
+ Qed.
+
+ Lemma c2tl_idem : forall (Gamma:list formula)(c:constant)(x:term),
+ c_appears_l Gamma c = false ->
+ c2tl Gamma c x = Gamma.
+ Proof.
+ induction Gamma.
+ simpl.
+ auto.
+ simpl.
+ intros.
+ apply orb_false_elim in H.
+ destruct H as [H1 H2].
+ rewrite IHGamma;auto.
+ rewrite c2t_idem; auto.
+ Qed.
+
+ Lemma c2t_lem2 : forall (A:formula)(c:constant)(s:term)(t:term),
+ c_appears_term t c = false -> loclt s ->
+ (c2t A c s) ^^ t = c2t (A ^^ t) c s.
+ Proof.
+ unfold open.
+ generalize 0.
+ fix 2.
+ intros.
+ destruct A.
+ simpl.
+ auto.
+ simpl.
+ rewrite c2t_lem2; auto.
+ rewrite c2t_lem2; auto.
+ simpl.
+ rewrite c2t_lem2; auto.
+ idtac.
+ induction t0. (* ili treba da bide t? *)
+ rewrite c2t_idem; auto.
+ rewrite c2t_idem; auto.
+ simpl.
+ destruct (beq_nat n n0).
+ assumption.
+ simpl.
+ reflexivity.
+ simpl.
+ reflexivity.
+ simpl.
+ unfold loclt in H0.
+ destruct (constant_dec c0 c).
+ rewrite H0.
+ reflexivity.
+ simpl.
+ reflexivity.
+ simpl.
+ simpl in IHt0.
+ assert (Hr :
+ (open_rec_term n t (c2t_term t0 c s)) =
+ (c2t_term (open_rec_term n t t0) c s)).
+ congruence.
+ rewrite Hr.
+ reflexivity.
+ Qed.
+
+ Lemma c2t_lem3 : forall (n : nat) (A : formula) (c : constant) (s t : term),
+ loclt s ->
+ open_rec n (c2t_term t c s) (c2t A c s) = c2t (open_rec n t A) c s.
+ Proof.
+ fix 2.
+ intros.
+ destruct A.
+ simpl.
+ auto.
+ simpl.
+ rewrite c2t_lem3.
+ rewrite c2t_lem3.
+ reflexivity.
+ assumption.
+ assumption.
+ simpl.
+ rewrite c2t_lem3; auto.
+ induction t0.
+ simpl.
+ destruct (beq_nat n n0).
+ reflexivity.
+ simpl.
+ reflexivity.
+ simpl.
+ reflexivity.
+ simpl.
+ destruct (constant_dec c0 c).
+ unfold loclt in H.
+ rewrite H.
+ reflexivity.
+ simpl.
+ reflexivity.
+ simpl.
+ assert ((open_rec_term n (c2t_term t c s) (c2t_term t0 c s))
+ = (c2t_term (open_rec_term n t t0) c s)).
+ simpl in IHt0.
+ congruence.
+ rewrite H0.
+ reflexivity.
+ Qed.
+
+ Lemma openrec_lem1 : forall A n t x,
+ open_rec n t (open_rec n (fvar x) A) =
+ open_rec n (fvar x) A.
+ Proof.
+ induction A; simpl; intros.
+
+ reflexivity.
+
+ rewrite IHA1.
+ rewrite IHA2.
+ reflexivity.
+
+ rewrite IHA.
+ reflexivity.
+
+ induction t.
+
+ simpl.
+ destruct (eq_nat_dec n n0).
+ rewrite <- e.
+ rewrite <- beq_nat_refl.
+ simpl.
+ reflexivity.
+ destruct beq_nat.
+ simpl; reflexivity.
+ simpl.
+ assert (beq_nat n n0 = false).
+ apply not_true_is_false.
+ intro H.
+ apply beq_nat_true in H.
+ congruence.
+ rewrite H.
+ reflexivity.
+
+ simpl.
+ reflexivity.
+
+ simpl.
+ reflexivity.
+
+ simpl.
+ congruence.
+ Qed.
+
+ Lemma locl_all_c2t : forall A,
+ locl A ->
+ forall c x, locl (c2t A c (fvar x)).
+ Proof.
+ unfold locl.
+ intros.
+ rewrite <- (H n (fvar x)) at 1.
+ rewrite <- c2t_lem3.
+ simpl c2t_term.
+ rewrite openrec_lem1.
+ rewrite <- (H n (fvar x)) at 2.
+ clear H.
+ generalize dependent n.
+ induction A; simpl; intros.
+
+ reflexivity.
+
+ rewrite IHA1.
+ rewrite IHA2.
+ reflexivity.
+
+ rewrite IHA.
+ reflexivity.
+
+ induction t0.
+
+ simpl.
+ destruct beq_nat.
+ simpl.
+ reflexivity.
+ simpl.
+ reflexivity.
+
+ simpl.
+ reflexivity.
+
+ simpl.
+ destruct constant_dec.
+ simpl.
+ reflexivity.
+ simpl.
+ reflexivity.
+
+ simpl.
+ congruence.
+
+ red.
+ simpl.
+ auto.
+ Qed.
+
+ Lemma thm283 : forall Gamma f x k,
+ proof Gamma f ->
+ proof (c2tl Gamma (added k) (fvar x)) (c2t f (added k) (fvar x)).
+ Proof.
+ intros Gamma f x k H.
+ induction H; set (c:=(added k)) in *; simpl; intros.
+
+ apply bot_elim.
+ simpl c2t in IHproof.
+ auto.
+
+ simpl c2t in IHproof1.
+ apply imp_elim with (c2t A c (fvar x)).
+ apply IHproof1.
+ auto.
+
+ simpl.
+ apply imp_intro.
+ simpl in IHproof.
+ auto.
+
+ simpl in IHproof.
+ apply dneg.
+ auto.
+
+ apply hypo.
+ induction Gamma.
+ simpl in *.
+ auto.
+ simpl in *.
+ destruct H.
+ rewrite H.
+ left; auto.
+ right; auto.
+
+ simpl in *.
+ unfold open.
+ rewrite <- c2t_lem3.
+ apply all_elim.
+ auto.
+ red.
+ reflexivity.
+
+ apply all_intro with L.
+ replace (all (c2t A c (fvar x))) with (c2t (all A) c (fvar x)).
+ apply locl_all_c2t; assumption.
+ reflexivity.
+ intros y yL.
+ rewrite c2t_lem2.
+ auto.
+ simpl.
+ reflexivity.
+ red; reflexivity.
+
+ apply weaken.
+ assumption.
+ Qed.
+
+ Lemma c2t_term_idem : forall c t, c2t_term (cnst c) c t = t.
+ intros.
+ simpl.
+ destruct constant_dec.
+ auto.
+ congruence.
+ Qed.
+End vanDalen_thm283.
+
+(** A section implementing the drinker paradox and another lemma, both
+ needed in henkin_equiconsistent. *)
+Section drinker.
+
+ (** some things are more naturally stated in terms of the
+ existential quantifier *)
+ Notation "'ex' x" := (neg (all (neg x))) (at level 0).
+
+ Lemma ex_intro : forall Delta t f, proof ((f^^t)::Delta) (ex f).
+ Proof.
+ intros.
+ impi.
+ impe (f^^t).
+ replace (neg (f^^t)) with ((neg f)^^t).
+ apply all_elim.
+ hypo.
+ reflexivity.
+ weak; hypo.
+ Qed.
+
+ Lemma ex_elim : forall Gamma f g, locl (all f) ->
+ proof Gamma (neg (all (neg f))) ->
+ (forall t, proof ((f^^t)::Gamma) g) ->
+ proof Gamma g.
+ Proof.
+ intros.
+ apply dneg.
+ apply imp_intro.
+ apply imp_elim with (all (neg f)).
+ weak.
+ assumption.
+ apply all_intro with nil.
+ apply locl_all_neg; assumption.
+ intros x xL.
+ unfold open.
+ simpl.
+ impi.
+ impe g.
+ weak; hypo.
+ apply weaken_lem1 with ((f^x)::Gamma).
+ red; simpl; intuition.
+ auto.
+ Qed.
+
+ Lemma lemma_HE1 : forall Delta h, locl (all (h ==> all h)) ->
+ proof Delta ((all (neg (h ==> all h))) ==> (neg ex (h ==> all h))).
+ Proof.
+ intros.
+ impi.
+ impi.
+ apply ex_elim with (h ==> all h).
+ assumption.
+ hypo.
+ intros.
+ impe ((h==>all h)^^t).
+ weak; weak.
+ replace (neg ((h ==> all h) ^^ t)) with ((neg (h ==> all h)) ^^ t).
+ apply all_elim.
+ hypo.
+ reflexivity.
+ hypo.
+ Qed.
+
+ (** and we also introduce disjunction *)
+ Notation "x \/ y" := ((neg x) ==> y).
+
+ Lemma disj_intro1 : forall Gamma f g, proof (f::Gamma) (f \/ g).
+ Proof.
+ intros.
+ impi.
+ apply bot_elim.
+ impe f.
+ hypo.
+ weak; hypo.
+ Qed.
+
+ Lemma disj_intro2 : forall Gamma f g, proof (g::Gamma) (f \/ g).
+ Proof.
+ intros.
+ impi.
+ weak; hypo.
+ Qed.
+
+ Lemma disj_elim : forall Gamma f g h,
+ proof (f::Gamma) h -> proof (g::Gamma) h -> proof Gamma ((f \/ g) ==> h).
+ Proof.
+ intros.
+ impi.
+ impe ((neg h ==> h)).
+ apply peirce.
+ apply proof_imp_trans with f.
+ apply proof_imp_trans with (neg g).
+ apply proof_imp_contrapos.
+ weak; impi; assumption.
+ impi.
+ dneg.
+ impi.
+ impe g.
+ weak; hypo.
+ impe (neg f).
+ weak; weak; hypo.
+ hypo.
+ weak; impi; assumption.
+ Qed.
+
+ Lemma LEM : forall Gamma f, proof Gamma (f \/ neg f).
+ Proof.
+ intros.
+ impi.
+ hypo.
+ Qed.
+
+ Lemma lemma_HE2_1 : forall h Delta, locl (all h) ->
+ proof (all h::Delta) ex (h ==> all h).
+ Proof.
+ intros.
+ pose (t := fvar 1).
+ set (f := h ==> all h).
+ impe (f ^^ t).
+ impi.
+ apply ex_intro.
+ unfold f.
+ assert (Hr:forall t, ((h ==> all h) ^^ t) = ((h^^t) ==> all h)).
+ unfold locl in H.
+ unfold open.
+ simpl in *.
+ intro t0.
+ rewrite H.
+ reflexivity.
+ rewrite Hr.
+ impi; weak; hypo.
+ Qed.
+
+ Lemma lemma_HE2_2 : forall h Delta, locl (all h) ->
+ proof (neg all h :: Delta) ex (neg h).
+ Proof.
+ intros.
+ impi.
+ impe (all h).
+ weak; hypo.
+ apply all_intro with nil.
+ assumption.
+ intros x xL.
+ dneg.
+ replace (neg neg (h^x)) with ((neg neg h)^x).
+ apply all_elim.
+ hypo.
+ reflexivity.
+ Qed.
+
+ (** Drinker paradox *)
+ Lemma lemma_HE2 : forall Delta h, locl (all h) ->
+ proof Delta (ex (h ==> all h)).
+ Proof.
+ intros.
+ impe (all h \/ neg all h).
+ Focus 2.
+ apply (LEM Delta (all h)).
+ apply disj_elim.
+ apply lemma_HE2_1.
+ assumption.
+ apply ex_elim with (neg h).
+ apply locl_all_neg; assumption.
+ apply lemma_HE2_2.
+ assumption.
+ intro t.
+ impe ((h==>all h)^^t).
+ impi.
+ apply ex_intro.
+ assert (Hr : forall t, ((h==>all h)^^t)=((h^^t)==>all h)).
+ intros.
+ replace ((h ==> all h) ^^ t0) with ((h^^t0) ==> ((all h)^^t0)).
+ unfold locl in H.
+ unfold open.
+ rewrite H.
+ reflexivity.
+ auto.
+
+ rewrite Hr.
+ impi.
+ apply bot_elim.
+ assert (Hr' : forall t, ((h==>bot)^^t)=((h^^t)==>bot)).
+ reflexivity.
+ rewrite Hr'.
+ impe (h^^t).
+ weak; hypo.
+ hypo.
+ Qed.
+End drinker.
+
+(** A predicate for distinguishing Henkin axioms *)
+Definition HA := fun f:formula =>
+ exists g:formula, f = ((g^^(cnst (added (all g)))) ==> all g)
+ /\ locl (all g).
+
+(** We also need a technical lemma for removing duplicate henkin axioms
+ from a context *)
+Section removing_duplicate_henkin_axioms.
+
+ Lemma rem_dup_lem1 :
+ forall a b g h, g<>h -> a = ((g ^^ cnst (added (all g))) ==> all g) ->
+ b = ((h ^^ cnst (added (all h))) ==> all h) ->
+ depth h <= depth g ->
+ c_appears b (added (all g)) = false.
+ Proof.
+ intros.
+ assert (c_appears h (added (all g)) = false).
+ apply bb'''.
+ simpl.
+ auto.
+ clear - H H1 H3.
+ rewrite H1.
+ simpl.
+ rewrite H3.
+ rewrite orb_false_r.
+ clear - H H3.
+ assert ((added (all h)) <> (added (all g))).
+ congruence.
+ clear H.
+ generalize dependent H0.
+ generalize ((added (all h))).
+ intros c Hc.
+ unfold open.
+ generalize 0.
+ induction h.
+ simpl in *.
+ reflexivity.
+ simpl in * |- .
+ apply orb_false_elim in H3.
+ destruct H3.
+ assert (IH1 := IHh1 H).
+ assert (IH2 := IHh2 H0).
+ simpl.
+ clear - IH1 IH2.
+ intro n.
+ rewrite IH1.
+ rewrite IH2.
+ auto.
+ simpl in *; unfold open in *.
+ intro n.
+ apply IHh.
+ assumption.
+ induction t.
+ Focus 2.
+ simpl in *.
+ reflexivity.
+ Focus 2.
+ simpl in *.
+ auto.
+ Focus 2.
+ simpl in *.
+ auto.
+ simpl in *.
+ intro m.
+ destruct beq_nat.
+ simpl.
+ destruct constant_dec.
+ congruence.
+ reflexivity.
+ simpl.
+ reflexivity.
+ Qed.
+
+ (** an order of deeper-than for the indexes of the Henkin
+ constants of two Henkin axioms; for non-Henkin axioms, no
+ behaviour is specified intentionally, but that's not a problem
+ because we use this order only for comparing Henkin axioms *)
+ Definition hgt (a b:formula) : bool :=
+ match a with
+ | (imp _ a') =>
+ match a' with
+ | (all a'') =>
+ match b with
+ | (imp _ b') =>
+ match b' with
+ | (all b'') => Compare_dec.leb (depth b'') (depth a'')
+ | _ => false
+ end
+ | _ => false
+ end
+ | _ => false
+ end
+ | _ => false
+ end.
+
+ Lemma rem_dup_HA : forall Gamma, included Gamma HA ->
+ exists Gamma', incl Gamma Gamma' /\ incl Gamma' Gamma /\
+ forall a Delta, Suffix (a::Delta) Gamma' ->
+ forall g, a = ((g ^^ cnst (added (all g))) ==> all g) ->
+ c_appears_l Delta (added (all g)) = false.
+ Proof.
+ intros.
+ assert (H0 := nodup_correct hgt formula_dec Gamma).
+ assert (H1 := sort_correct hgt (nodup hgt formula_dec Gamma)).
+ set (Gamma' := sort hgt (nodup hgt formula_dec Gamma)) in *.
+ exists Gamma'.
+ destruct H0.
+ destruct H2.
+ destruct H1.
+ destruct H4.
+ intuition.
+ eauto using incl_tran.
+ eauto using incl_tran.
+ assert (H8:=H5 a Delta H6).
+ assert (H9:=Suffix_incl H6).
+ assert (H10:forall b, In b Delta -> a<>b).
+ apply Suffix_lem1 in H6.
+ destruct H6 as [l Hl].
+ assert (NoDup Gamma').
+ unfold Gamma'.
+ apply nodup_sort.
+ assumption.
+ rewrite <- Hl in H6.
+ apply NoDup_remove_2 in H6.
+ intros b Hb Heq.
+ rewrite <- Heq in Hb.
+ clear -H6 Hb.
+ assert (In a (l++Delta)).
+ apply in_or_app.
+ auto.
+ auto.
+ clear H5 H6.
+ induction Delta.
+ simpl.
+ reflexivity.
+ simpl.
+ rewrite IHDelta.
+ rewrite orb_false_r.
+ assert (HA a0).
+ apply H.
+ apply H2.
+ apply H4.
+ clear -H9.
+ red in H9.
+ auto using in_eq,in_cons.
+ destruct H5 as [h [Hh Hhclosed]].
+ apply rem_dup_lem1 with a h.
+ assert (a<>a0).
+ apply H10.
+ simpl.
+ auto.
+ rewrite Hh in H5.
+ rewrite H7 in H5.
+ clear -H5.
+ congruence.
+ assumption.
+ assumption.
+ assert (hgt a a0 = true).
+ apply H8.
+ simpl; auto.
+ rewrite Hh in H5.
+ rewrite H7 in H5.
+ clear -H5.
+ simpl in H5.
+ apply leb_complete.
+ assumption.
+ clear -H8.
+ intros.
+ apply H8.
+ simpl.
+ auto.
+ clear -H9.
+ unfold incl in *.
+ intros.
+ apply H9.
+ simpl in H |- *.
+ intuition.
+ clear -H10.
+ intros.
+ apply H10.
+ simpl.
+ auto.
+ Qed.
+End removing_duplicate_henkin_axioms.
+
+Module CBAproof_completion := filter_completion CBAproof.
+Export CBAproof_completion.
+
+Section Completeness.
+
+ (** T1 is an extension of T2, both sets of formulas *)
+ Definition extension (T1 T2:formula_set) := forall f, T2 f -> T1 f.
+
+ (** A classical theory is a set of formulas T closed under
+ derivation over a set of Axioms *)
+ Definition theory (Axioms:formula_set)(T:formula_set) :=
+ forall f:formula,
+ (T f <->
+ exists Gamma:list formula,
+ (included Gamma Axioms /\ exists p:proof Gamma f, True)).
+
+ (** A set of formulas is Henkin-complete if it contains a Henkin
+ axiom for any locally closed formula. *)
+ Definition henkin_complete (T:formula_set) :=
+ forall f:formula,
+ locl (all f) ->
+ T ((f^^(cnst (added (all f)))) ==> all f).
+
+ (** Two sets of formulas being equiconsistent *)
+ Definition equicons (X Y:formula_set) := X bot <-> Y bot.
+
+ (** An axiom set is extended with Henkin axioms *)
+ Definition AxH (T A:formula_set) := fun f:formula => A f \/ HA f.
+
+ (** A theory that closes the extended axiom set under derivation *)
+ Definition TH (T A:formula_set) := fun f:formula =>
+ exists Gamma:list formula,
+ included Gamma (AxH T A) /\ exists p:proof Gamma f, True.
+
+ (** When an axiom set contains no added constants *)
+ Definition noHC (A:formula_set) := forall f, A f -> added_cnsts_f f = false.
+
+ Notation "'ex' x" := (neg (all (neg x))) (at level 0).
+ Open Scope type_scope.
+
+ (** If T is a theory over an axiom set which contains no Henkin
+ constants, then (TH T) is equiconsistent with T. *)
+ Lemma henkin_equiconsistent : forall A T:formula_set, noHC A -> theory A T ->
+ (TH T A) bot -> T bot.
+ Proof.
+ intros A T HcleanA H1 H2.
+ destruct H2 as [Gamma H3].
+ destruct H3.
+ destruct H0 as [p _].
+ (* We want to show that if Eta contains a henkin axiom, it can
+ be eliminated from the proof. *)
+ assert (lemma1:forall h Delta Eta, locl (all h) ->
+ c_appears_l Delta (added (all h)) = false ->
+ Eta = ((h ^^ cnst (added (all h))) ==> all h)::Delta ->
+ proof Eta bot ->
+ proof Delta bot
+ ).
+ clear.
+ intros h Delta Eta closed_h Hc_Delta Hr p.
+ rewrite Hr in p.
+ clear Hr.
+ apply imp_intro in p.
+ assert (H : proof Delta (all (neg (h ==> all h)))).
+ apply all_intro with nil.
+ apply locl_henkin_neg; assumption.
+ intros x xL.
+ replace (neg ((h ^^ cnst (added (all h))) ==> all h))
+ with ((neg (h ==> all h)) ^^ cnst (added (all h))) in p.
+ assert (p' := thm283 x (all h) p).
+ rewrite c2tl_idem in p'; auto.
+ unfold open in p'.
+ rewrite <- c2t_lem3 in p'; auto.
+ rewrite c2t_term_idem in p'.
+ rewrite c2t_idem in p'.
+ assumption.
+ clear.
+ simpl.
+ rewrite c_appears_thm.
+ auto.
+ red; reflexivity.
+ transitivity (neg ((h ^^ cnst (added (all h))) ==> ((all h) ^^ cnst (added (all h))))).
+ reflexivity.
+ unfold locl in closed_h.
+ unfold open.
+ rewrite closed_h.
+ reflexivity.
+ assert (H0 : proof Delta ((neg (ex (h ==> all h))))).
+ impe (all (neg (h ==> all h))).
+ apply lemma_HE1.
+ apply locl_henkin; assumption.
+ assumption.
+ impe (ex (h ==> all h)).
+ apply H0.
+ clear H.
+ apply lemma_HE2.
+ assumption.
+ (* end of proof of lemma1 *)
+ assert (lemma3 : forall Eta, included Eta (AxH T A) ->
+ exists E1, exists E2, incl Eta (E1++E2) /\
+ included E1 HA /\ included E2 A).
+ clear.
+ intros.
+ induction Eta.
+ exists nil.
+ exists nil.
+ simpl.
+ unfold incl.
+ auto using nil_included.
+ assert (H1 : included Eta (AxH T A)).
+ clear -H.
+ unfold included in *.
+ intros.
+ auto using in_cons.
+ destruct (IHEta H1) as [E1 [E2 [H2 [H3 H4]]]].
+ assert (H5 : (AxH T A) a).
+ red in H.
+ apply H.
+ apply in_eq.
+ unfold AxH in H5.
+ destruct H5.
+ exists E1.
+ exists (a::E2).
+ intuition.
+ clear -H2.
+ red in H2 |- *.
+ intros.
+ simpl in H.
+ destruct H.
+ rewrite <- H.
+ clear.
+ induction E1.
+ simpl.
+ auto.
+ simpl.
+ auto.
+ assert (H3 := H2 _ H).
+ clear -H3.
+ induction E1.
+ simpl in *; auto.
+ simpl in *.
+ destruct H3; auto.
+ clear - H4 H0.
+ red.
+ intros.
+ simpl in H.
+ destruct H.
+ rewrite <- H.
+ assumption.
+ red in H4.
+ auto.
+ exists (a::E1).
+ exists E2.
+ intuition.
+ simpl.
+ clear H5.
+ red.
+ simpl.
+ intros.
+ intuition.
+ clear -H3 H0.
+ red.
+ red in H3.
+ intros.
+ simpl in H.
+ destruct H.
+ rewrite <- H.
+ assumption.
+ apply H3.
+ assumption.
+ (* end of proof of lemma3 *)
+ assert (lemma2 : forall Eta, included Eta (AxH T A) ->
+ proof Eta bot -> exists Delta, included Delta A * proof Delta bot).
+ clear - lemma1 lemma3 HcleanA.
+ intros.
+ assert (H3 := lemma3 Eta H).
+ clear lemma3 H.
+ destruct H3 as [E1 [E2 [HEta [HE1 HE2]]]].
+ exists E2.
+ intuition.
+ (* begin proof using rem_dup_HA *)
+ destruct (rem_dup_HA HE1) as [E1' HE1'].
+ destruct HE1' as [HE1E1' [HE1'E1 HE1']].
+ assert (proof (E1++E2) bot).
+ apply weaken_lem1 with Eta; auto.
+ assert (proof (E1'++E2) bot).
+ apply weaken_lem1 with (E1++E2); auto.
+ clear -HE1E1'.
+ unfold incl in *.
+ intros.
+ apply in_or_app.
+ apply in_app_or in H.
+ intuition.
+ clear -H1 HE1' lemma1 HE1'E1 HE1 HcleanA HE2.
+ induction E1'.
+ simpl in H1.
+ assumption.
+
+ apply IHE1'.
+ clear -HE1'E1.
+ unfold incl in *.
+ intros.
+ apply HE1'E1.
+ apply in_cons.
+ assumption.
+ intros.
+ assert (Suffix (a0::Delta) (a::E1')).
+ constructor.
+ assumption.
+ assert (H3 := HE1' a0 Delta H2 g H0).
+ assumption.
+ rewrite <- app_comm_cons in H1.
+ assert (HA a).
+ clear -HE1 HE1'E1.
+ unfold included in HE1.
+ apply HE1.
+ red in HE1'E1.
+ apply HE1'E1.
+ simpl.
+ auto.
+ unfold HA in H.
+ destruct H as [g [Hg Hgclosed]].
+ apply lemma1 with g (a::(E1'++E2)).
+ assumption.
+ Focus 2.
+ rewrite Hg; reflexivity.
+ Focus 2.
+ assumption.
+
+ assert (H2 := HE1' a E1' (Suffix_refl _) g Hg).
+ clear - HcleanA H2 HE2.
+ unfold noHC in HcleanA.
+ (* znaci tuka uste ni treba hipotezata deka E2 ne sodrzi henkin konstanti *)
+ apply c_appears_l_app.
+ intuition.
+ clear - HE2 HcleanA.
+ induction E2.
+ simpl.
+ reflexivity.
+ simpl.
+ apply orb_false_intro.
+ apply added_cnsts_c_appears'.
+ unfold included in HE2.
+ simpl in *.
+ apply HcleanA.
+ apply HE2.
+ auto.
+ apply IHE2.
+ clear - HE2.
+ unfold included in *.
+ simpl in *.
+ intros.
+ auto.
+(* end proof using rem_dup_HA *)
+ assert (H2 := lemma2 Gamma H p).
+ clear lemma2.
+ apply (proj2 (H1 bot)).
+ destruct H2 as [Delta [H2 H3]].
+ exists Delta.
+ intuition.
+ exists H3.
+ constructor.
+ Qed.
+
+ (** If T is a theory whose axiom set has no Henkin constants, then
+ (TH T) is a theory which is Henkin-complete and equiconsistent
+ with T. *)
+ Theorem enrich : forall T A:formula_set, noHC A -> theory A T ->
+ extension (AxH T A) A /\ extension (TH T A) T /\
+ theory (AxH T A) (TH T A) /\ henkin_complete (TH T A) /\ equicons (TH T A) T.
+ Proof.
+ intros T A AnoHC TAtheory.
+ (* (TH T A) is an extension of T *)
+ assert (extAxH : extension (AxH T A) A).
+ unfold extension.
+ red.
+ intuition.
+ assert (extTHT : extension (TH T A) T).
+ unfold extension.
+ red.
+ intros f Tf.
+ unfold theory in TAtheory.
+ assert (H1 := proj1 (TAtheory f) Tf).
+ destruct H1 as [Gamma HGamma].
+ exists Gamma.
+ destruct HGamma as [HG1 HG2].
+ split.
+ unfold included in HG1 |- *.
+ unfold extension in extAxH.
+ auto.
+ exact HG2.
+ intuition.
+ (* (TH T A) is a theory with axiom set (AxH T A) *)
+ unfold theory.
+ intros f.
+ split.
+ intro HT.
+ destruct HT as [Gamma HGamma].
+ exists Gamma.
+ split.
+ intuition.
+ intuition.
+ intros.
+ unfold TH.
+ exact H.
+ (* (TH T A) is henkin_complete *)
+ unfold henkin_complete.
+ intros f.
+ exists (((f ^^ cnst (added (all f))) ==> all f)::nil).
+ split.
+ red.
+ unfold AxH.
+ intros.
+ right.
+ unfold HA.
+ exists f.
+ simpl in H0.
+ intuition.
+ assert (p : proof (((f ^^ cnst (added (all f))) ==> all f) :: nil)
+ ((f ^^ cnst (added (all f))) ==> all f)).
+ hypo.
+ exists p.
+ auto.
+ (* (TH T A) and T are equiconsistent *)
+ red.
+ split; intros.
+ apply (@henkin_equiconsistent A T AnoHC TAtheory).
+ assumption.
+ intuition.
+ Qed.
+
+ (** A theory is also a filter *)
+ Theorem theory2filter : forall T Ax, theory Ax T -> Filter T.
+ Proof.
+ simpl.
+ intros T Ax HT.
+ unfold formula_set in *.
+ unfold theory in *.
+ constructor.
+ (* a theory is nonempty, it contains top *)
+ exists top.
+ simpl.
+ assert (H := proj2 (HT (neg bot))).
+ apply H.
+ clear HT H.
+ exists (@nil formula).
+ split.
+ apply nil_included.
+ assert (proof nil (neg bot)).
+ apply imp_intro.
+ apply hypo.
+ simpl.
+ auto.
+ exists H; auto.
+ (* a theory is closed under the ordering induced by the boolean algebra *)
+ intros.
+ red in H0.
+ destruct H0.
+ clear H0.
+ destruct H1.
+ assert (H' := proj1 (HT (x)) H).
+ destruct H'.
+ destruct H1.
+ destruct H2.
+ generalize dependent (x).
+ generalize dependent (y).
+ clear -HT H1.
+ intros Y X FX H2 H3.
+ assert (HT' := HT Y).
+ assert (HT'' := proj2 HT').
+ clear HT HT'.
+ apply HT''.
+ clear HT''.
+ exists x1.
+ split.
+ assumption.
+ assert (proof x1 Y).
+ dneg.
+ impi.
+ impe (X ==> neg Y).
+ impe X.
+ apply weaken_lem1 with (@nil formula).
+ apply nil_lem1.
+ assumption.
+ weak;assumption.
+ impi.
+ weak;hypo.
+ exists H;auto.
+ (* a theory is closed under finite meets *)
+ intros.
+ assert (H1 := proj1 (HT (x)) H).
+ assert (H2 := proj1 (HT (y)) H0).
+ destruct H1.
+ destruct H2.
+ simpl.
+ generalize dependent (x).
+ generalize dependent (y).
+ clear -HT.
+ intros Y FY H1 X FX H2.
+ assert (HT' := HT (neg (X ==> neg Y))).
+ assert (HT'' := proj2 HT').
+ apply HT''.
+ clear HT HT' HT''.
+ exists (x0++x1).
+ split.
+ firstorder using included_lem1 || (* 8.1 *) firstorder with included_lem1.
+ destruct H1 as [H21 H22].
+ destruct H2 as [H51 H52].
+ destruct H22 as [p1 _].
+ destruct H52 as [p2 _].
+ assert (proof (x0++x1) (neg (X ==> neg Y))).
+ clear -p1 p2.
+ impi.
+ impe Y.
+ impe X.
+ hypo.
+ weak.
+ apply weaken_lem1 with x0.
+ apply incl_appl.
+ apply incl_refl.
+ assumption.
+ weak.
+ apply weaken_lem1 with x1.
+ apply incl_appr.
+ apply incl_refl.
+ assumption.
+ exists H; auto.
+ Qed.
+
+ (** A subsection which implements the model existence lemma by using
+ [enrich] and the ultrafilter Z from filters.v *)
+ Section model_existence.
+ Variables (T Ax:formula_set)(AxnoHC:noHC Ax)(Ttheory:theory Ax T).
+
+ (** A Henkin extension of T *)
+ Definition T1 : formula -> Prop := TH T Ax.
+ Definition Ax1 : formula -> Prop := AxH T Ax.
+ Definition T1theory := (proj1 (proj2 (proj2 (enrich AxnoHC Ttheory)))).
+ Definition T1filter : Filter T1 := theory2filter T1theory.
+
+ (** Extend T1 to a meta-dn filter using the Z defined in filters.v *)
+ Definition Z' := Z T1.
+
+ (** The properties of Z, proved in thm22 in filters.v *)
+ Definition lem15 := thm22 T1filter.
+
+ Lemma model_existence1 : extension Z' T /\ equicons T Z'.
+ Proof.
+ intros.
+ assert (Hext : extension Z' T).
+ unfold Z', Z in *.
+ unfold extension.
+ intros.
+ exists 0.
+ simpl.
+ unfold T1.
+ simpl.
+ unfold TH.
+ destruct (Ttheory f).
+ clear H1.
+ destruct (H0 H).
+ destruct H1.
+ exists x.
+ intuition.
+ clear -H1.
+ unfold AxH.
+ unfold included in *.
+ intros.
+ left.
+ auto.
+ intuition.
+ (* equiconsistency: *)
+ red.
+ set (HT1 := enrich AxnoHC Ttheory).
+ destruct HT1.
+ destruct H0.
+ destruct H1.
+ destruct H2.
+ assert (HZ:=lem15).
+ destruct HZ.
+ destruct H5.
+ red in H5.
+ unfold inconsistent in H5.
+ intuition.
+ apply (@henkin_equiconsistent Ax T).
+ assumption.
+ assumption.
+ unfold T1 in *.
+ assumption.
+ Qed.
+
+ (** An abbreviation for F_n which lives in formula->Prop *)
+ Definition G := F_ T1.
+
+ (** The accompanying axioms *)
+ Fixpoint GAx (n':nat) : formula -> Prop :=
+ match n' with
+ | O => Ax1
+ | S n =>
+ let Fn := F_ T1 n
+ in fun f:formula =>
+ GAx n f \/
+ enum f = n /\
+ equiconsistent Fn (up (union_singl Fn f))
+ end.
+
+ (* The system of axioms for the complete filter Z *)
+ Definition ZAx := fun f:formula => exists n:nat, GAx n f.
+
+ Lemma GAx_monotone : forall n m:nat, le n m ->
+ forall f, GAx n f -> GAx m f.
+ Proof.
+ intros.
+ induction H.
+ assumption.
+ simpl.
+ left.
+ assumption.
+ Qed.
+
+ Lemma remove_In_neq_In : forall ys y xn,
+ In y ys -> y <> xn -> In y (remove id_B_dec xn ys).
+ Proof.
+ induction ys; intros.
+ simpl in *.
+ contradiction.
+ simpl in *.
+ destruct H.
+ destruct id_B_dec.
+ congruence.
+ simpl.
+ intuition.
+ destruct id_B_dec.
+ auto.
+ simpl.
+ right.
+ auto.
+ Qed.
+
+ Lemma proof_fold_left : forall Gamma f,
+ proof Gamma f -> fold_left meet Gamma top <= f.
+ Proof.
+ intro Gamma.
+ induction Gamma.
+ simpl.
+ intros.
+ unfold leq.
+ unfold eq_B.
+ split.
+ assert (p: proof nil (top && f ==> top)).
+ unfold meet.
+ unfold top.
+ impi.
+ impi.
+ hypo.
+ exists p;constructor.
+ assert (p:proof nil (top ==> top&&f)).
+ unfold meet.
+ impi.
+ impi.
+ impe f.
+ impe top.
+ hypo.
+ weak;hypo.
+ weak;weak;assumption.
+ exists p;constructor.
+ intros.
+ apply imp_intro in H.
+ assert (IH := IHGamma _ H).
+ clear -IH.
+ rewrite fold_left_meet_cons.
+ set (phi := fold_left meet Gamma top) in *.
+ unfold leq in *.
+ unfold eq_B in *.
+ split.
+ assert (forall a b, proof nil (meet a b ==> a)).
+ clear.
+ intros.
+ unfold meet.
+ impi.
+ dneg.
+ impi.
+ impe (a ==> neg b).
+ weak; hypo.
+ impi.
+ apply bot_elim.
+ impe a.
+ weak;hypo.
+ hypo.
+ assert (p : proof nil (a && phi && f ==> a && phi)).
+ apply H.
+ exists p.
+ constructor.
+ destruct IH as [[H1 _] [H2 _]].
+ assert (proof nil (a && phi ==> a && phi && f)).
+ assert (forall a b Delta, proof Delta (a ==> b) -> proof Delta (a ==> meet a b)).
+ intros.
+ unfold meet.
+ impi.
+ impi.
+ impe b.
+ impe a0.
+ hypo.
+ weak;hypo.
+ impe a0.
+ weak;weak;assumption.
+ weak;hypo.
+ apply H.
+ clear -H2.
+ unfold meet in *.
+ impi.
+ dneg.
+ impi.
+ impe (a ==> neg phi).
+ weak; hypo.
+ impi.
+ impi.
+ impe (phi ==> neg (a ==> f)).
+ apply weaken_lem1 with (phi::nil).
+ red.
+ simpl.
+ intuition.
+ impe phi.
+ weak; assumption.
+ hypo.
+ impi.
+ impi.
+ impe f.
+ weak;weak;weak;weak;hypo.
+ impe a.
+ hypo.
+ weak;weak;weak;hypo.
+ exists H.
+ constructor.
+ Qed.
+
+ Lemma fold_left_proof : forall Gamma f,
+ fold_left meet Gamma top <= f -> proof Gamma f.
+ Proof.
+ intro Gamma.
+ induction Gamma.
+ simpl.
+ intros f H.
+ unfold leq,eq_B in H.
+ destruct H as [[p1 _] [p2 _]].
+ clear -p2.
+ unfold meet in *.
+ dneg.
+ impi.
+ impe (top ==> neg f).
+ impe top.
+ weak; assumption.
+ impi; hypo.
+ impi;weak;hypo.
+ intros f H.
+ rewrite fold_left_meet_cons in H.
+ assert (IH:=IHGamma (imp a f)).
+ impe a.
+ Focus 2.
+ hypo.
+ weak.
+ apply IH.
+ clear -H.
+ set (phi := fold_left meet Gamma top) in *.
+ unfold leq,eq_B in H |- *.
+ destruct H as [[p1 _] [p2 _]].
+ split.
+ clear.
+ assert (p:proof nil (phi && (a ==> f) ==> phi)).
+ unfold meet.
+ impi.
+ dneg.
+ impi.
+ impe (phi==>neg(a==>f)).
+ weak;hypo.
+ impi.
+ apply bot_elim.
+ impe phi.
+ weak;hypo.
+ hypo.
+ exists p;constructor.
+ assert (p:proof nil (phi ==> phi && (a ==> f))).
+ impi.
+ impi.
+ impe (a==>f).
+ impe phi.
+ hypo.
+ weak;hypo.
+ impi.
+ impe (a&&phi&&f).
+ impi.
+ unfold meet.
+ dneg.
+ impi.
+ impe (neg (a==>neg phi)==>neg f).
+ weak;hypo.
+ impi;weak;hypo.
+ impe (a&&phi).
+ weak;weak;weak;assumption.
+ impi.
+ impe phi.
+ impe a.
+ hypo.
+ weak;hypo.
+ weak;weak;weak;hypo.
+ exists p;constructor.
+ Qed.
+
+ (** Every filter F_n is also a theory *)
+ Lemma F_n_theory : forall n, theory (GAx n) (G n).
+ Proof.
+ induction n.
+ (* base case *)
+ unfold G.
+ simpl.
+ constructor.
+ intros.
+ assumption.
+ intros.
+ assert (H2 := proj2 (T1theory f)).
+ apply H2.
+ clear H2.
+ assumption.
+
+ (* induction case *)
+ constructor.
+
+ intro H.
+ assert (exists Gamma,
+ included Gamma (GAx (S n)) /\
+ fold_left meet Gamma top <= f
+ ).
+ simpl in H.
+ destruct H as [m [ys [H1 [H2 H3]]]].
+ assert (forall y, In y ys -> exists zs, included zs (GAx (S n)) /\
+ fold_left meet zs top <= y).
+ intros y yys.
+ assert (G n y \/ GAx (S n) y).
+ assert (H4 := lemma5 _ _ _ H2).
+ case H4.
+ intro case1.
+ left.
+ clear -yys case1.
+ induction ys.
+ simpl in yys.
+ contradiction.
+ simpl in case1.
+ apply lemma61 in case1.
+ simpl in yys.
+ case yys.
+ intro case11.
+ rewrite <- case11.
+ intuition.
+ destruct case1.
+ intuition.
+ intro case2.
+ clear -yys case2.
+ destruct case2 as (xn & xnys & nxn & Hfold & Hequi).
+ case (formula_dec y xn).
+ intro case21.
+ right.
+ rewrite case21.
+ symmetry in nxn.
+ simpl.
+ intuition.
+ intro case22.
+ left.
+ set (ys' := remove id_B_dec xn ys) in *.
+ assert (yys' : In y ys').
+ apply remove_In_neq_In; assumption.
+ clear -yys' Hfold.
+ induction ys'.
+ simpl in yys'.
+ contradiction.
+ simpl in *.
+ apply lemma61 in Hfold.
+ case yys'.
+ intro case221.
+ rewrite <- case221.
+ intuition.
+ intro case222.
+ intuition.
+ (* *)
+ case H.
+ intro case1.
+ assert (Hth := proj1 (IHn y) case1).
+ clear -Hth.
+ simpl.
+ destruct Hth as [Gamma [HGamma1 [HGamma2 _]]].
+ exists Gamma.
+ unfold included in *.
+ intuition.
+ apply proof_fold_left.
+ assumption.
+ intro case2.
+ exists (y::nil).
+ split.
+ unfold included.
+ simpl In.
+ intros yy Hr.
+ destruct Hr.
+ rewrite <- H0.
+ assumption.
+ contradiction.
+ simpl.
+ rewrite meet_top.
+ apply leq_refl.
+ (* *)
+ clear -H H3 IHn.
+ assert (exists Gamma : list formula,
+ included Gamma (GAx (S n)) /\ fold_left meet Gamma top <= fold_left meet ys top).
+ clear H3.
+ generalize dependent ys.
+ induction ys.
+ intros.
+ simpl fold_left.
+ clear H.
+ exists nil.
+ split.
+ apply nil_included.
+ simpl.
+ apply leq_refl.
+ intros.
+ assert (H1 : forall y, In y ys ->
+ exists zs : list formula,
+ included zs (GAx (S n)) /\ fold_left meet zs top <= y).
+ clear -H.
+ simpl In in H.
+ intros.
+ intuition.
+ assert (H2 :
+ exists zs : list formula,
+ included zs (GAx (S n)) /\ fold_left meet zs top <= a).
+ clear -H.
+ simpl In in H.
+ intros.
+ intuition.
+ assert (IH := IHys H1).
+ clear IHys H H1.
+ destruct H2 as (zs & Hzs1 & Hzs2).
+ destruct IH as (Gamma & HG1 & HG2).
+ exists (zs++Gamma).
+ split.
+ auto using included_lem1.
+ clear -Hzs2 HG2.
+ rewrite fold_left_meet_cons.
+ rewrite fold_left_app_lem.
+ apply meet_leq_compat; assumption.
+ (* *)
+ destruct H0 as [Gamma [HG1 HG2]].
+ exists Gamma.
+ intuition.
+ apply leq_trans with (fold_left meet ys top).
+ assumption.
+ assumption.
+ destruct H0 as (Gamma & HG1 & HG2).
+ exists Gamma.
+ intuition.
+ clear -HG2.
+ assert (p := fold_left_proof _ HG2).
+ exists p.
+ constructor.
+ (* second direction *)
+ intro H.
+ simpl in *.
+ unfold up at 1.
+ destruct H as [Gamma [H01 [H02 _]]].
+ exists (length Gamma).
+ exists Gamma.
+ intuition.
+ clear H02.
+ induction Gamma.
+ simpl in *.
+ constructor.
+ simpl.
+ apply lemma62 ; auto.
+ split.
+ apply IHGamma.
+ clear -H01.
+ unfold included in *.
+ simpl in *.
+ auto.
+ clear - H01 IHn.
+ unfold included in *.
+ simpl in *.
+ destruct H01 with a.
+ auto.
+ left.
+ apply (proj2 (IHn a)).
+ exists (a::nil).
+ split.
+ red.
+ simpl.
+ clear -H.
+ intros.
+ intuition.
+ rewrite <- H1.
+ assumption.
+ assert (p:proof (a::nil) a).
+ hypo.
+ exists p; auto.
+ right.
+ assumption.
+ (* now the first subcase *)
+ apply proof_fold_left; assumption.
+ Qed.
+
+ Theorem Z'theory : theory ZAx Z'.
+ Proof.
+ constructor.
+ intro H0.
+ unfold Z' in H0.
+ destruct H0 as [n Hn].
+ assert (H1 := proj1 (F_n_theory n f)).
+ unfold G in H1.
+ assert (H3 := H1 Hn).
+ clear -H3.
+ destruct H3 as [Gamma [H1 H2]].
+ exists Gamma.
+ intuition.
+ clear -H1.
+ unfold included in *.
+ unfold ZAx.
+ intros.
+ exists n; auto.
+ (* the other direction left (if it's necessary) *)
+ intro.
+ unfold Z'.
+ unfold Z.
+ destruct H as [Gamma [H1 [H2 _]]].
+ unfold ZAx in H1.
+ (* Gamma is a finite list, so there is a maximum n, such that
+ the entire list belongs to (GAx n) *)
+ assert (exists m:nat, included Gamma (GAx m)).
+ clear -H1.
+ unfold included in *.
+ induction Gamma.
+ simpl in *.
+ exists 0.
+ intuition.
+ assert ((forall f : formula, In f Gamma -> exists n : nat, GAx n f)).
+ clear -H1.
+ simpl in *.
+ intros.
+ intuition.
+ assert (IH := IHGamma H).
+ clear H IHGamma.
+ destruct IH as [m Hm].
+ assert (H1' := H1 a).
+ clear H1.
+ simpl in *.
+ assert (exists n:nat, GAx n a).
+ intuition.
+ clear H1'.
+ destruct H as [n Hn].
+ destruct (Compare_dec.le_lt_dec n m).
+ assert (H1:=GAx_monotone l).
+ exists m.
+ intuition.
+ rewrite <- H0 in *.
+ auto.
+ assert (le m n).
+ apply Lt.lt_le_weak; assumption.
+ assert (H1:=GAx_monotone H).
+ exists n.
+ intuition.
+ rewrite <- H2 in *.
+ auto.
+ (* *)
+ destruct H as [m Hm].
+ exists m.
+ assert (H3:=proj2 ((F_n_theory m) f)).
+ apply H3.
+ exists Gamma.
+ intuition.
+ exists H2.
+ constructor.
+ Qed.
+
+ Definition metaDN (X:formula_set) := forall f:formula,
+ (X (neg f) -> X bot) -> X f.
+
+ Lemma metaDNZ': metaDN Z'.
+ Proof.
+ destruct lem15 as [_ [_ H]].
+ unfold metaDN.
+ intros.
+ unfold complete in H.
+ unfold Z'.
+ assert (H3 := H f).
+ clear H.
+ unfold element_complete in H3.
+ apply H3.
+ clear H3.
+ unfold equiconsistent.
+ split.
+ unfold inconsistent.
+ intro.
+ apply lem223; auto.
+ unfold union_singl.
+ left; assumption.
+ (* the harder direction: *)
+ intro.
+ unfold inconsistent.
+ unfold Z' in H0.
+ apply H0.
+ clear H0.
+ unfold inconsistent in H.
+ destruct H as [m Hm].
+ destruct Hm as [ys [H3 [H4 H5]]].
+ simpl in H5.
+ assert (H6 := lemma8 (Z T1) f ys H4).
+ clear H4.
+ destruct H6.
+ Focus 2.
+ set (z':=(remove id_B_dec f ys)) in *.
+ set (z:=fold_left meet z' top) in *.
+ assert ((meet f z) <= bott).
+ apply leq_trans with (fold_left meet ys top); auto.
+ apply leq_trans with (fold_left meet (f::z') top); auto.
+ apply eq_B_leq; auto.
+ simpl.
+ symmetry.
+ unfold z.
+ apply fold_left_meet_cons; auto.
+ apply lemma2; auto.
+ unfold z'.
+ apply lemma3; auto.
+ clear -H0 Ttheory H.
+ assert (z <= neg f).
+ destruct H0 as [[p1 _] [p2 _]].
+ unfold leq.
+ unfold eq_B.
+ unfold meet.
+ split.
+ assert (proof nil (neg (z ==> neg neg f) ==> z)).
+ impi.
+ dneg.
+ impi.
+ impe (z==>neg neg f).
+ weak;hypo.
+ impi.
+ apply bot_elim.
+ impe z.
+ weak;hypo.
+ hypo.
+ exists H0; auto.
+ assert (proof nil (z ==> neg (z ==> neg neg f))).
+ unfold meet in p2.
+ impi.
+ impi.
+ impe (neg f).
+ impe z.
+ hypo.
+ weak;hypo.
+ weak.
+ impi.
+ impe (neg (f ==> neg z) ==> neg bott).
+ impe (neg (f ==> neg z)).
+ weak;weak;assumption.
+ impi.
+ impe z.
+ impe f.
+ hypo.
+ weak;hypo.
+ weak;weak;hypo.
+ impi.
+ weak;weak;weak.
+ impi.
+ hypo.
+ exists H0;auto.
+ apply (upwards_closed (proj1 lem15)) with z; auto.
+ unfold z.
+ apply lemma4; auto.
+ apply (proj1 lem15).
+ idtac.
+ assert (bott <= neg f).
+ apply leq_bott;auto.
+ apply (upwards_closed (proj1 lem15)) with bott; auto.
+ apply (upwards_closed (proj1 lem15)) with (fold_left meet ys top); auto.
+ apply lemma4; auto.
+ apply (proj1 lem15).
+ Qed.
+
+ Lemma model_existence2 :
+ (forall A B : formula, Z' (A ==> B) -> Z' A -> Z' B) /\
+ (forall A B : formula, (Z' A -> Z' B) -> Z' (A ==> B)).
+ Proof.
+ (*
+ Z is a theory, because it's a filter, so direction => by modus
+ ponens. Direction <=:
+
+ Spse Z A implies Z B. To show: Z (A==>B).
+
+ As Z is complete, Meta-DN holds i.e. for all closed C, ((Z ~C)
+ -> Z bot) -> Z C. So, we have to show that ((Z ~(A==>B)) -> Z
+ bot). From Z ~(A==>B) we get both Z A and Z ~B. From Z A and
+ the hypothesis we get Z B, hence Z bot.
+ *)
+ assert (dir1: forall A B0 : formula, Z' (A ==> B0) -> Z' A -> Z' B0).
+ intros.
+ apply (proj2 (Z'theory B0)).
+ assert (H1':= proj1 (Z'theory _) H).
+ assert (H2':= proj1 (Z'theory _) H0).
+ clear -H1' H2'.
+ destruct H1' as [Gamma1 [H11 H12]].
+ destruct H2' as [Gamma2 [H21 H22]].
+ exists (Gamma1++Gamma2).
+ split.
+ red.
+ intros.
+ destruct (in_app_or Gamma1 Gamma2 f H); auto.
+ assert (proof (Gamma1 ++ Gamma2) B0).
+ destruct H12 as [p1 _].
+ destruct H22 as [p2 _].
+ clear -p1 p2.
+ impe A.
+ apply weaken_lem1 with Gamma1.
+ apply incl_appl.
+ apply incl_refl.
+ assumption.
+ apply weaken_lem1 with Gamma2.
+ apply incl_appr.
+ apply incl_refl.
+ assumption.
+ exists H; auto.
+ split.
+ apply dir1.
+ (* direction <= *)
+ intros.
+ apply metaDNZ'.
+ intro.
+ apply dir1 with B0.
+ assert (H2':= proj1 (Z'theory _) H0).
+ destruct H2' as [Gamma2 [H21 H22]].
+ destruct H22 as [p2 _].
+ assert (proof Gamma2 (neg B0)).
+ impi.
+ impe (A ==> B0).
+ weak; assumption.
+ impi.
+ weak; hypo.
+ apply (proj2 (Z'theory (neg B0))).
+ exists Gamma2.
+ split.
+ assumption.
+ exists H1; auto.
+ apply H.
+ assert (H2':= proj1 (Z'theory _) H0).
+ destruct H2' as [Gamma2 [H21 H22]].
+ destruct H22 as [p2 _].
+ assert (proof Gamma2 A).
+ dneg.
+ impi.
+ impe (A ==> B0).
+ weak; assumption.
+ impi.
+ apply bot_elim.
+ impe A.
+ weak; hypo.
+ hypo.
+ apply (proj2 (Z'theory A)).
+ exists Gamma2.
+ split.
+ assumption.
+ exists H1; auto.
+ Qed.
+
+ Lemma model_existence3' : henkin_complete Z'.
+ Proof.
+ unfold henkin_complete.
+ intros.
+ unfold Z'.
+ assert (T1 ((f ^^ cnst (added (all f))) ==> all f)).
+ unfold T1.
+ unfold TH.
+ exists (((f ^^ cnst (added (all f))) ==> all f) :: nil).
+ split.
+ unfold AxH.
+ red.
+ intros.
+ right.
+ red.
+ simpl in H0.
+ intuition.
+ exists f.
+ intuition.
+ (* this was repeating a part of the proof of 'enrich' *)
+ assert (p : proof (((f ^^ cnst (added (all f))) ==> all f) :: nil)
+ ((f ^^ cnst (added (all f))) ==> all f)).
+ hypo.
+ exists p; auto.
+ unfold Z.
+ exists 0.
+ simpl.
+ assumption.
+ Qed.
+
+ Lemma model_existence3 :
+ (forall A, Z' (all A) -> (forall t:term, Z' (A^^t))) /\
+ (forall A, locl (all A) -> (forall t:term, loclt t -> Z' (A^^t)) -> Z' (all A)).
+ Proof.
+ assert (Z'imp := proj1 model_existence2).
+ split.
+ (* first part: Z' is a filter and so a theory. QED by all_elim. *)
+ intros.
+ apply (proj2 (Z'theory (A^^t))).
+ assert (H0' := proj1 (Z'theory _) H).
+ clear H.
+ destruct H0' as [Gamma [H01 H02]].
+ destruct H02 as [H02 _].
+ assert (proof Gamma (A ^^ t)).
+ apply all_elim.
+ assumption.
+ exists Gamma.
+ intuition.
+ exists H; auto.
+ (* second part: use henkin completeness *)
+ intros.
+ apply Z'imp with (A ^^ cnst (added (all A))).
+ apply model_existence3'.
+ assumption.
+ assert (H2 : forall t:term, loclt t ->
+ exists Gamma : list formula,
+ included Gamma ZAx /\ exists H : proof Gamma (A^^t), True).
+ intros t Hct.
+ apply (proj1 (Z'theory (A^^t))).
+ auto.
+ apply H0.
+ red.
+ simpl.
+ auto.
+ Qed.
+
+ Theorem model_existence : extension Z' T /\ model Z' /\ equicons T Z'.
+ Proof.
+ assert (H1 := model_existence1).
+ intuition.
+ constructor.
+ intros.
+ apply (proj2 (Z'theory (neg neg A ==> A))).
+ exists (@nil formula).
+ split.
+ apply nil_included.
+ assert (proof nil (neg neg A ==> A)).
+ impi.
+ dneg.
+ hypo.
+ exists H1; auto.
+ apply (proj1 (model_existence2)).
+ apply (proj2 (model_existence2)).
+ apply (proj1 (model_existence3)).
+ apply (proj2 (model_existence3)).
+ Qed.
+
+ End model_existence.
+
+ Fixpoint HCl (A:list formula) : bool :=
+ match A with
+ | nil => false
+ | cons f fs => orb (added_cnsts_f f) (HCl fs)
+ end.
+
+ Lemma HCl_correct : forall f Gamma, In f Gamma -> HCl Gamma = false ->
+ added_cnsts_f f = false.
+ Proof.
+ intros.
+ induction Gamma.
+ contradict H.
+ simpl in H0.
+ apply orb_false_elim in H0.
+ destruct H0.
+ simpl in H.
+ destruct H.
+ rewrite <- H; assumption.
+ auto.
+ Qed.
+
+ Definition noHCf (f:formula) := added_cnsts_f f = false.
+ Definition noHCl (A:list formula) := HCl A = false.
+
+ (** Finally, the completeness theorem. If Gamma and A contain no Henkin constants, that is, if they are built up of/like usual formulas, and if A is true in every model in which Gamma is true, then there is a derivation of A with hypothesis form Gamma. *)
+ Theorem completeness : forall Gamma A, noHCl Gamma -> noHCf A ->
+ valid Gamma A -> proof Gamma A.
+ Proof.
+ intros Gamma A Hno1 Hno2 H.
+ set (Ax := fun f:formula => In f Gamma \/ f = neg A).
+ set (T := fun f:formula =>
+ exists Delta:list formula,
+ included Delta Ax /\ exists p:proof Delta f, True).
+ assert (noHCAx : noHC Ax).
+ clear - Hno1 Hno2.
+ unfold noHC,noHCl,noHCf in *.
+ intros f H.
+ unfold Ax in H.
+ destruct H.
+ eauto using HCl_correct.
+ rewrite H.
+ simpl.
+ rewrite Hno2.
+ auto.
+ assert (Ttheory : theory Ax T).
+ unfold theory.
+ unfold T.
+ intuition.
+ assert (H1 := model_existence noHCAx Ttheory).
+ destruct H1 as [Mext [Mmodel Mequicons]].
+ set (M:=Z' T Ax).
+ dneg.
+ impi.
+ assert (T bot).
+ assert (M bot).
+ apply model_imp_faithful1 with A.
+ assumption.
+ assert (T (neg A)).
+ red.
+ exists ((neg A)::nil).
+ split.
+ red.
+ intros.
+ simpl in H0.
+ red.
+ right.
+ intuition.
+ assert (proof ((neg A) :: nil) (neg A)).
+ hypo.
+ exists H0.
+ auto.
+ unfold M.
+ intuition.
+ red in H.
+ assert (H' := H M Mmodel).
+ red in H'.
+ apply H'.
+ intros.
+ assert (T f).
+ red.
+ exists Gamma.
+ split.
+ red.
+ red.
+ intuition.
+ assert (proof Gamma f).
+ hypo.
+ exists H1.
+ auto.
+ unfold M.
+ intuition.
+ red in Mequicons.
+ intuition.
+ red in H0.
+ destruct H0 as [Delta [H01 H02]].
+ destruct H02 as [p _].
+ apply weaken_lem1 with Delta.
+ clear - H01.
+ red.
+ intros.
+ red in H01.
+ simpl.
+ case (H01 a H).
+ intuition.
+ intuition.
+ assumption.
+ Qed.
+
+End Completeness.
+
+End classical_completeness.
diff --git a/boolean_completeness/filters.v b/boolean_completeness/filters.v
new file mode 100644
index 0000000..2f90a06
--- /dev/null
+++ b/boolean_completeness/filters.v
@@ -0,0 +1,962 @@
+(** A proof that every filter F over a countable boolean algebra can
+ be extended to a complete filter equiconsistent with
+ F. (Ultrafilter Theorem)
+
+ This file is used by the completeness theorem from classcomp.v, but
+ it can be used in other contexts, it is a standalone module.
+ *)
+Require Import Ring.
+Require Import List.
+Require Import Compare_dec.
+Require Import Setoid.
+Require Import Morphisms.
+
+Set Implicit Arguments.
+
+(** Definition of Countable Boolean Algebra (over a setoid) *)
+Module Type CBA.
+ Parameter Inline B:Set.
+ Parameters (meet join:B -> B -> B)(bott top:B)(compl:B -> B).
+ Parameter eq_B : B -> B -> Prop.
+
+ Notation "x == y" := (eq_B x y) (at level 70, no associativity).
+ Axiom eq_B_refl : reflexive B eq_B.
+ Axiom eq_B_symm : symmetric B eq_B.
+ Axiom eq_B_trans : transitive B eq_B.
+
+ Axiom eq_B_meet_morph : Proper (eq_B ==> eq_B ==> eq_B) meet.
+ Axiom eq_B_join_morph : Proper (eq_B ==> eq_B ==> eq_B) join.
+
+ Axiom enum : B -> nat.
+ Axiom countable : forall x y, enum x = enum y -> x = y.
+
+ Axiom meet_idem : forall x, meet x x == x.
+ Axiom join_idem : forall x, join x x == x.
+ Axiom meet_comm : forall x y, meet x y == meet y x.
+ Axiom join_comm : forall x y, join x y == join y x.
+ Axiom meet_assoc : forall x y z, meet x (meet y z) == meet (meet x y) z.
+ Axiom join_assoc : forall x y z, join x (join y z) == join (join x y) z.
+ Axiom meet_absorb : forall x y, meet x (join x y) == x.
+ Axiom join_absorb : forall x y, join x (meet x y) == x.
+ Axiom meet_distrib : forall x y z,
+ meet (join x y) z == join (meet x z) (meet y z).
+ Axiom join_distrib : forall x y z,
+ join (meet x y) z == meet (join x z) (join y z).
+ Axiom meet_bott : forall x, meet bott x == bott.
+ Axiom join_bott : forall x, join bott x == x.
+ Axiom meet_top : forall x, meet top x == x.
+ Axiom join_top : forall x, join top x == top.
+ Axiom meet_compl : forall x, meet x (compl x) == bott.
+ Axiom join_compl : forall x, join x (compl x) == top.
+
+ (** We also need that identity [id_B] is decidable in 1 place. Note that this is not that [id_B] is definitional equality, it has nothing to do with the equality of the setoid. *)
+ Axiom id_B_dec : forall x y : B, {x = y}+{x <> y}.
+ Axiom id_B_dec_right : forall (x y:B), x<>y ->
+ exists H:x<>y, id_B_dec x y = right (x=y) H.
+ Axiom id_B_dec_left : forall x:B,
+ exists H:x=x, id_B_dec x x = left (x<>x) H.
+
+End CBA.
+
+Module filter_completion (cba : CBA).
+ Export cba.
+
+ Add Relation B eq_B
+ reflexivity proved by eq_B_refl
+ symmetry proved by eq_B_symm
+ transitivity proved by eq_B_trans
+ as eq_B_relation.
+
+ Add Morphism meet with signature eq_B ==> eq_B ==> eq_B as meet_morphism.
+ exact eq_B_meet_morph.
+ Qed.
+
+ Add Morphism join with signature eq_B ==> eq_B ==> eq_B as join_morphism.
+ exact eq_B_join_morph.
+ Qed.
+
+ (** A boolean algebra is also a semi-ring, useful because Coq can automatically solve equations for in such cases. *)
+ Theorem CBA_semiring : semi_ring_theory bott top join meet eq_B.
+ Proof.
+ constructor.
+ apply join_bott.
+ apply join_comm.
+ apply join_assoc.
+ apply meet_top.
+ apply meet_bott.
+ apply meet_comm.
+ apply meet_assoc.
+ apply meet_distrib.
+ Qed.
+
+ Add Ring B_semiring : CBA_semiring.
+
+ (** Lattice as an algebra is equivalent to a lattice as a poset, so
+ we can define an ordering. *)
+ Definition leq := fun x y => meet x y == x.
+
+ Lemma leq' : forall x y, leq x y <-> join x y == y.
+ Proof.
+ intros.
+ intuition.
+ unfold leq in *.
+ rewrite <- H .
+ rewrite meet_comm.
+ rewrite join_comm.
+ apply join_absorb.
+ unfold leq.
+ rewrite <- H .
+ apply meet_absorb.
+ Qed.
+
+ Delimit Scope B_scope with B.
+ Bind Scope B_scope with B.
+ Arguments Scope B [nat_scope].
+ Open Scope B_scope.
+
+ Notation "x && y" := (meet x y) (at level 40, left associativity) : B_scope.
+ Notation "- x" := (compl x) (at level 35, right associativity) : B_scope.
+ Notation "x || y" := (join x y) (at level 50, left associativity) : B_scope.
+ Notation "x <= y" := (leq x y) : B_scope.
+
+ (** When a subset F of B is a filter *)
+ Record Filter (F:B -> Prop) : Prop := {
+ nonempty : exists x:B, F x;
+ upwards_closed : forall x y:B, F x -> leq x y -> F y;
+ meet_closed : forall x y:B, F x -> F y -> F (meet x y)
+ }.
+
+ (** The closure of the subset X, [up X], is the least filter containing X.
+
+ Conjuction of finite (but arbitrary) lenght is represented by the
+ List operation fold_left. *)
+ Definition up (X:B -> Prop) := fun z:B =>
+ exists n:nat, exists ys:list B, length ys = n /\
+ fold_left (fun (a:Prop)(b:B) => and a (X b)) ys True /\
+ leq (fold_left meet ys top) z.
+
+ Definition union_singl (F:B -> Prop)(x:B) := fun b:B => F b \/ b=x.
+
+ Definition inconsistent (F:B -> Prop) := F bott.
+
+ Definition equiconsistent (F G:B -> Prop) := inconsistent F <-> inconsistent G.
+
+ Definition element_complete (F:B -> Prop)(x:B) :=
+ equiconsistent F (up (union_singl F x)) -> F x.
+
+ (** This notion of completeness of a filter is the key to having a constructive proof. This notion is constructivelly weaker than, but classically equivalent to, the more usual notion: either x in F, or ~x in F. *)
+ Definition complete (F:B -> Prop) := forall x:B, element_complete F x.
+
+ Lemma leq_refl : forall x:B, x <= x.
+ Proof.
+ intro x.
+ unfold leq.
+ apply meet_idem.
+ Qed.
+
+ Lemma leq_trans : forall x y z:B, leq x y -> leq y z -> leq x z.
+ Proof.
+ unfold leq.
+ intros.
+ rewrite <- H.
+ ring [H0].
+ Qed.
+
+ Lemma eq_B_leq : forall x y:B, x==y -> x <= y.
+ Proof.
+ intros x y.
+ unfold leq.
+ intro H0.
+ rewrite H0.
+ apply meet_idem.
+ Qed.
+
+ Add Morphism leq with signature eq_B ==> eq_B ==> iff as leq_morphism.
+ Proof.
+ intros x y H x0 y0 H0.
+ split.
+ intro.
+ apply leq_trans with x.
+ apply eq_B_leq.
+ symmetry.
+ assumption.
+ apply leq_trans with x0.
+ assumption.
+ apply eq_B_leq.
+ assumption.
+ intro.
+ apply leq_trans with y.
+ apply eq_B_leq.
+ assumption.
+ apply leq_trans with y0.
+ assumption.
+ apply eq_B_leq.
+ symmetry.
+ assumption.
+ Qed.
+
+ Lemma meet_leq_compat : forall a b c d, a <= b -> c <= d -> a && c <= b && d.
+ Proof.
+ unfold leq.
+ intros.
+ ring [H H0].
+ Qed.
+
+ (** The next few lemmas with names "fold_left*" and "lemma*" are used to handle the representation of a finite number of conjunctions by a list. *)
+
+ Lemma fold_left_meet_morphism : forall x y bs, x==y ->
+ fold_left meet bs x == fold_left meet bs y.
+ Proof.
+ intros.
+ generalize dependent x.
+ generalize dependent y.
+ induction bs.
+ auto.
+ simpl.
+ intros.
+ apply IHbs.
+ ring [H].
+ Qed.
+
+ Lemma fold_left_meet_cons : forall bs b, fold_left meet (cons b bs) top == b && (fold_left meet bs top).
+ Proof.
+ simpl.
+ induction bs.
+ simpl.
+ intro.
+ ring.
+ intro.
+ simpl.
+ rewrite IHbs.
+ rewrite meet_assoc.
+ transitivity (fold_left meet bs (top && (b && a))).
+ apply fold_left_meet_morphism.
+ rewrite <- meet_assoc.
+ reflexivity.
+ apply IHbs.
+ Qed.
+
+ Lemma fold_left_impl : forall (X:B -> Prop)(xs:list B)(Q P:Prop), (Q -> P) -> fold_left (fun (a : Prop) (b : B) => a /\ X b) xs Q -> fold_left (fun (a : Prop) (b : B) => a /\ X b) xs P.
+ Proof.
+ induction xs.
+ simpl.
+ auto.
+ simpl.
+ intros.
+ apply IHxs with (Q /\ X a); intuition.
+ Qed.
+
+ Lemma fold_left_app_lem : forall xs ys,
+ fold_left meet (xs ++ ys) top ==
+ fold_left meet xs top && fold_left meet ys top.
+ Proof.
+ intros xs ys.
+ induction xs.
+ simpl.
+ ring.
+ rewrite <- app_comm_cons.
+ rewrite fold_left_meet_cons.
+ rewrite fold_left_meet_cons.
+ rewrite <- meet_assoc.
+ rewrite IHxs.
+ reflexivity.
+ Qed.
+
+ Lemma up_filter : forall X:B -> Prop, Filter (up X).
+ Proof.
+ intro X.
+ constructor.
+
+ (* up X is nonempty, because it contains at least top *)
+ exists top.
+ unfold up.
+ exists 0.
+ exists (@nil B).
+ simpl.
+ intuition.
+ apply leq_refl.
+
+ (* up X is upwards_closed, because if x \in up X and x <= y then we
+ can use the same sequence of witnesses for x's membership to X
+ and the transitivity of <= to conclude that y \in up X. *)
+ intros.
+ destruct H as [n [ys [H1 [H2 H3]]]].
+ exists n.
+ exists ys.
+ intuition.
+ apply leq_trans with x; auto.
+
+ (* up X is meet-closed, by lemma meet_leq_compat. *)
+ intros.
+ unfold up in *.
+ destruct H as [n [xs [H1 [H2 H3]]]].
+ destruct H0 as [m [ys [H4 [H5 H6]]]].
+ exists (n+m).
+ exists (xs++ys).
+ rewrite app_length.
+ rewrite fold_left_app.
+ intuition.
+ apply fold_left_impl with True; auto.
+
+ rewrite fold_left_app_lem.
+ apply meet_leq_compat; assumption.
+ Qed.
+
+ Lemma filter_top : forall F:B -> Prop, Filter F -> F top.
+ Proof.
+ intros.
+ destruct (nonempty H) as [w Fw].
+ apply upwards_closed with w; auto.
+ unfold leq.
+ ring.
+ Qed.
+
+ Lemma lemma3 : forall (T:Type)(Hdec:forall x y : T, {x = y} + {x <> y})(x:T)(ys:list T),
+ incl ys (x :: remove Hdec x ys).
+ Proof.
+ intros.
+ induction ys.
+ simpl.
+ unfold incl.
+ intros.
+ absurd (In a nil).
+ apply in_nil.
+ assumption.
+ simpl.
+ unfold incl in *.
+ intros.
+ assert (H' := in_inv H).
+ case H'.
+ intro Hr.
+ rewrite <- Hr.
+ case (Hdec x a).
+ intro.
+ rewrite e.
+ apply in_eq.
+ intro.
+ apply in_cons.
+ apply in_eq.
+ intro.
+ assert (IH:=IHys _ H0).
+ case (Hdec x a).
+ intro.
+ assumption.
+ intro.
+ simpl.
+ simpl in IH.
+ intuition.
+ Qed.
+
+ Lemma lemma2 : forall A C:list B, incl A C ->
+ leq (fold_left meet C top) (fold_left meet A top).
+ Proof.
+ induction A.
+
+ simpl.
+ intros.
+ unfold leq.
+ ring.
+
+ intros.
+ rewrite fold_left_meet_cons.
+ assert (IHA' := IHA C).
+
+ (* a is in C, so foldleft C = a && foldleft C, by one rewrite <- meet_idem *)
+
+ apply leq_trans with (a && fold_left meet C top).
+ unfold incl in H.
+ assert (H' := H a (in_eq a A)).
+ clear -H'.
+ induction C.
+ absurd (In a nil); auto.
+ simpl In in H'.
+ case H'.
+ intro Ha.
+ rewrite Ha.
+ rewrite fold_left_meet_cons.
+ rewrite meet_assoc.
+ rewrite meet_idem.
+ apply leq_refl.
+ intro Ca.
+ rewrite fold_left_meet_cons.
+ rewrite meet_assoc.
+ rewrite (meet_comm a a0).
+ rewrite <- meet_assoc.
+ apply meet_leq_compat.
+ apply leq_refl.
+ auto.
+ apply meet_leq_compat.
+ apply leq_refl.
+ apply IHA'.
+ apply incl_tran with (a::A).
+ apply incl_tl.
+ apply incl_refl.
+ assumption.
+ Qed.
+
+ (* membership in a filter is also a morphism, but how to declare as such? *)
+ Lemma filter_memb_morph : forall F, Filter F ->
+ forall (x y:B), x==y -> F x -> F y.
+ Proof.
+ intros.
+ apply upwards_closed with x; auto.
+ apply eq_B_leq.
+ assumption.
+ Qed.
+
+ Lemma lemma4 : forall xs F, Filter F ->
+ fold_left (fun (A : Prop) (b : B) => A /\ F b) xs True ->
+ F (fold_left meet xs top).
+ Proof.
+ intros.
+ induction xs.
+ simpl in *.
+ apply (@filter_top F).
+ assumption.
+ apply filter_memb_morph with (a && fold_left meet xs top); auto.
+ rewrite fold_left_meet_cons.
+ reflexivity.
+ simpl in H0.
+ apply (@meet_closed F).
+ assumption.
+ induction xs.
+ simpl in *.
+ intuition.
+ apply IHxs0.
+ simpl in H0.
+ apply fold_left_impl with (((True /\ F a) /\ F a0)).
+ intuition.
+ assumption.
+ intro.
+ assert (Hr := fold_left_meet_cons xs a0).
+ apply upwards_closed with ((a0 && fold_left meet xs top)).
+ assumption.
+ apply filter_memb_morph with (fold_left meet (a0 :: xs) top).
+ assumption.
+ assumption.
+ apply IHxs.
+ apply fold_left_impl with ((True /\ F a)).
+ intuition.
+ assumption.
+ unfold leq.
+ rewrite <- meet_assoc.
+ rewrite meet_idem.
+ reflexivity.
+ apply IHxs.
+ apply fold_left_impl with ((True /\ F a)).
+ intuition.
+ assumption.
+ Qed.
+
+ Lemma lemma61 : forall (f:B -> Prop)(l:list B)(basecase:Prop)(P:Prop),
+ fold_left (fun (R:Prop)(x:B) => R /\ (f x)) l (basecase /\ P) ->
+ (fold_left (fun (R:Prop)(x:B) => R /\ (f x)) l basecase) /\ P.
+ Proof.
+ induction l.
+ simpl in *.
+ auto.
+ simpl.
+ clear -IHl.
+ intros.
+ apply IHl.
+ apply fold_left_impl with ((basecase /\ P) /\ f a).
+ intro.
+ intuition.
+ assumption.
+ Qed.
+
+ Lemma lemma62 : forall (f:B -> Prop)(l:list B)(basecase:Prop)(P:Prop),
+ (fold_left (fun (R:Prop)(x:B) => R /\ (f x)) l basecase) /\ P ->
+ fold_left (fun (R:Prop)(x:B) => R /\ (f x)) l (basecase /\ P).
+ Proof.
+ induction l.
+ simpl in *.
+ auto.
+ simpl.
+ clear -IHl.
+ intros.
+ apply fold_left_impl with ((basecase /\ f a) /\ P).
+ intro.
+ intuition.
+ apply IHl.
+ assumption.
+ Qed.
+
+ Lemma lemma5 : forall (X:B -> Prop)(ys:list B)(n:nat)
+ (H:fold_left (fun (P : Prop) (x : B) => P /\ (X x \/
+ (enum x = n /\ equiconsistent X (up (union_singl X x))))) ys True),
+ fold_left (fun (P : Prop) (x : B) => P /\ X x) ys True \/
+ (exists x_n : B,
+ In x_n ys /\
+ n = enum x_n /\
+ fold_left (fun (P : Prop) (x : B) => P /\ X x) (remove id_B_dec x_n ys) True /\
+ equiconsistent X (up (union_singl X x_n))
+ ).
+ Proof.
+ fix 2.
+ induction ys.
+ intros.
+ simpl in *.
+ left.
+ auto.
+ intros.
+ simpl in H.
+
+ assert (H' := lemma61 _ _ _ _ H).
+ destruct H' as [H'' case_a ].
+ assert (H5 := lemma5 _ _ _ H'').
+ clear - H5 case_a IHys.
+ simpl.
+ case H5.
+ (* case 1*)
+ intro H.
+ clear H5.
+ case case_a.
+ intro Xa.
+ left.
+ apply lemma62.
+ intuition.
+ intro H1.
+ right.
+ exists a.
+ split.
+ left.
+ reflexivity.
+ split.
+ symmetry.
+ intuition.
+ split; auto.
+ assert (Hr := id_B_dec_left a).
+ destruct Hr as [Hr1 Hr2].
+ rewrite Hr2.
+ clear -H.
+ induction ys.
+ simpl in *.
+ auto.
+ simpl in H.
+ assert (H' := lemma61 _ _ _ _ H).
+ destruct H' as [H'' Xa0].
+ assert (IH := IHys H'').
+ clear - Xa0 IH.
+(* now if a=a0 then (remove id_B_dec a (a0 :: ys))=(remove id_B_dec a ys) QED *)
+(* and if a<>a0 then (remove id_B_dec a (a0 :: ys))=a0::(remove id_B_dec a ys) *)
+(* QED by Xa0 and IH and lemma62 *)
+ simpl remove.
+ case (id_B_dec a a0).
+ auto.
+ intro.
+ simpl.
+ apply lemma62.
+ intuition.
+ intuition.
+ (* case 2 *)
+ intro H.
+ clear H5.
+ destruct H as [xn [H1 [H2 [H3 H4]]]].
+ (* subcase a \in X *)
+ case case_a.
+ intro Xa.
+ right.
+ exists xn.
+ split.
+ right.
+ assumption.
+ split; auto.
+ split; auto.
+ assert (Hdec := id_B_dec xn a).
+ case Hdec.
+ intro Hr.
+ rewrite <- Hr.
+ assert (Hleft := id_B_dec_left xn).
+ destruct Hleft as [Hleft1 Hleft2].
+ rewrite Hleft2.
+ assumption.
+ intro Hineq.
+ assert (Hright := id_B_dec_right Hineq).
+ destruct Hright as [Hright1 Hright2].
+ rewrite Hright2.
+ simpl.
+ apply lemma62.
+ split; auto.
+ (* subcase a = x_n *)
+ intro HH.
+ destruct HH as [Henum H5].
+ right.
+ exists a.
+ split.
+ left; reflexivity.
+ split.
+ symmetry; intuition.
+ assert (Hleft := id_B_dec_left a).
+ destruct Hleft as [Hleft1 Hleft2].
+ rewrite Hleft2.
+ assert (a=xn).
+ apply countable.
+ rewrite Henum.
+ rewrite <- H2.
+ reflexivity.
+ rewrite H.
+ intuition.
+ Qed.
+
+ Section completion.
+
+ (** This is the hearth of the argument, a fixpoint that, starting from a filter F, builds a complete filter extending F and equiconsistent to F. *)
+ Variable F:B -> Prop.
+
+ Fixpoint F_ (n':nat) {struct n'} :=
+ match n' with
+ | (S n) =>
+ let Fn := F_ n in
+ let X := fun b:B => Fn b \/
+ (enum b=n /\ equiconsistent Fn (up (union_singl Fn b)))
+ in up X
+ | O => F
+ end.
+
+ Definition Z := fun b:B => exists n:nat, (F_ n) b.
+
+ Lemma lem223 : forall (X:B -> Prop) x, X x -> (up X) x.
+ Proof.
+ intros.
+ unfold up.
+ exists 1.
+ exists (x::nil).
+ simpl.
+ intuition.
+ unfold leq.
+ rewrite meet_top.
+ rewrite meet_idem.
+ reflexivity.
+ Qed.
+
+ Lemma lem222 : forall n m, (n <= m)%nat -> forall x, (F_ n) x -> (F_ m) x.
+ Proof.
+ assert (lem2221 : forall k, forall x, (F_ k) x -> (F_ (S k)) x).
+ intros.
+ simpl.
+ apply lem223.
+ left; assumption.
+ intros.
+ (* induction on the relation <= *)
+ induction H.
+ assumption.
+ apply lem2221.
+ assumption.
+ Qed.
+
+ Theorem thm22 : Filter F ->
+ Filter Z /\ equiconsistent F Z /\ complete Z.
+ Proof.
+ intros.
+ assert (lem221 : forall n, Filter (F_ n)).
+ intro.
+ case n.
+ simpl.
+ assumption.
+ intro m.
+ simpl.
+ apply up_filter.
+ (* QED lem221 *)
+
+ assert (Fn_filter : forall n:nat, Filter (F_ n)).
+ intro.
+ case n.
+ simpl.
+ assumption.
+ intro.
+ simpl.
+ apply up_filter.
+ (* QED Fn_filter *)
+
+ (* Z is a filter
+ - Z is nonempty: it has the same witness of nonemptyness as F_0 *)
+ split.
+ constructor.
+ assert (H1:=nonempty H).
+ destruct H1 as [b Fb].
+ exists b.
+ exists 0.
+ simpl.
+ assumption.
+ (* - Z is upwards closed: if x \in Z then there is a n s.t. x \in
+ F_n. If nnow x <= y, since F_n is a filter, y \in F_n and so
+ y\in Z.*)
+ idtac.
+ intros.
+ destruct H0 as [n Fnx].
+ exists n.
+ apply upwards_closed with x; auto.
+ idtac.
+ (* - Z is closed under meets: if x,y \in Z, then there are n,m
+ such that x \in F_n, y \in F_m. We use the fact that any F_k
+ is a filter and a lemma saying that: k <= l -> F_k \subseteq
+ F_l. From the decidability of inequality on nat, we conclude
+ the proof. *)
+ intros.
+ destruct H0 as [n Fnx].
+ destruct H1 as [m Fmx].
+ case (le_lt_dec n m).
+ intro.
+ exists m.
+ apply (@meet_closed (F_ m)); auto.
+ apply lem222 with n;auto.
+ intro.
+ assert (l' := Lt.lt_le_weak _ _ l).
+ exists n.
+ apply (@meet_closed (F_ n)); auto.
+ apply lem222 with m;auto.
+
+ (* Some lemmas needed botth for the equiconsistency of F and Z and
+ for the completeness of Z *)
+ assert (lem224 : forall n, equiconsistent (F_ n) F).
+ induction n.
+ simpl.
+ unfold equiconsistent.
+ intuition.
+ split.
+ intro.
+ unfold inconsistent in *.
+ simpl F_ in H0.
+ destruct H0 as [l [ys [H11 [H12 H13]]]].
+ destruct IHn as [IH1 IH2].
+ apply IH1.
+ clear IH1 IH2.
+ assert (CaseAnalysis :
+ (fold_left (fun (a : Prop) (b : B) => a /\ F_ n b) ys True) \/
+ exists x_n,
+ In x_n ys /\
+ n = enum x_n /\
+ (fold_left (fun (a : Prop) (b : B) => a /\ F_ n b) (remove id_B_dec x_n ys) True) /\
+ equiconsistent (F_ n) (up (union_singl (F_ n) x_n))
+ ).
+ apply lemma5.
+ assumption.
+ case CaseAnalysis.
+ (* case 1 *)
+ unfold inconsistent.
+ intros H1.
+ apply upwards_closed with (fold_left meet ys top);auto.
+ clear - H1 H Fn_filter.
+ apply (@lemma4 ys (F_ n)).
+ auto.
+ assumption.
+ (* case 2 *)
+ intros H1.
+ clear - H1 H11 H12 H13 lem221.
+ destruct H1 as [x_n [H21 [H22 [H23 H24]]]].
+ set (Y := remove id_B_dec x_n ys) in *.
+ assert (A0 : leq (fold_left meet (x_n :: Y) top) bott).
+ apply leq_trans with (fold_left meet ys top); auto.
+ apply lemma2.
+ unfold Y.
+ apply lemma3.
+ rewrite fold_left_meet_cons in A0.
+ assert (A1 : leq (fold_left meet Y top) (compl x_n)).
+ set (y:=fold_left meet Y top) in *.
+ clear -A0.
+ unfold leq in *.
+ rewrite meet_comm in A0.
+ rewrite meet_bott in A0.
+ assert (A0' : bott || - x_n == x_n && y || - x_n).
+ rewrite A0.
+ reflexivity.
+ rewrite join_bott in A0'.
+ rewrite join_distrib in A0'.
+ rewrite join_compl in A0'.
+ rewrite meet_top in A0'.
+ symmetry in A0'.
+ assert (H := leq' y (- x_n)).
+ unfold leq in *.
+ intuition.
+ assert (A2 : F_ n (fold_left meet Y top)).
+ apply (@lemma4 Y (F_ n)).
+ apply lem221.
+ assumption.
+ assert (A3 : F_ n (compl x_n)).
+ apply upwards_closed with (fold_left meet Y top).
+ apply lem221.
+ assumption.
+ assumption.
+ clear - A3 H24 lem221 .
+ unfold equiconsistent in H24.
+ destruct H24 as [H241 H242].
+ apply H242.
+ clear - A3 lem221 .
+ unfold inconsistent.
+ assert (A4 : up (union_singl (F_ n) x_n) (- x_n)).
+ apply lem223.
+ unfold union_singl.
+ intuition.
+ assert (A5 : up (union_singl (F_ n) x_n) x_n).
+ apply lem223.
+ unfold union_singl.
+ intuition.
+ apply filter_memb_morph with (x_n && - x_n).
+ apply up_filter.
+ apply meet_compl.
+ apply (@meet_closed (up (union_singl (F_ n) x_n))).
+ apply up_filter.
+ assumption.
+ assumption.
+ (* end of cases 1 and 2 *)
+ unfold inconsistent.
+ simpl.
+ intro Fincon.
+ destruct IHn.
+ apply lem223.
+ left.
+ apply H1.
+ assumption.
+ (* QED lem224 *)
+ assert (lem225 : forall n m, equiconsistent (F_ n) (F_ m)).
+ intros n m.
+ clear - lem224.
+ firstorder.
+ (* QED lem225 *)
+ assert (lem226 : equiconsistent F Z).
+ unfold Z.
+ unfold equiconsistent.
+ unfold inconsistent.
+ split.
+ intro Fincon.
+ exists 0.
+ simpl.
+ assumption.
+ intro Zincon.
+ destruct Zincon.
+ clear -H0 lem224.
+ firstorder.
+ (* QED lem226 *)
+
+ split.
+ (* F and Z are equiconsistent *)
+ apply lem226.
+
+ (* one more lemma needed for "Z is complete" *)
+ assert (subseteq_up_mono : forall (X Y:B -> Prop)(b:B), (forall z, X z -> Y z) -> up X b -> up Y b).
+ unfold up.
+ intros.
+ destruct H1 as [n [ys [H1 [H2]]]].
+(* intro H3. *)
+ exists n.
+ exists ys.
+ intuition.
+ clear H1 H3.
+ induction ys.
+ simpl in *; auto.
+ simpl in H2 |- *.
+ assert (H2' := lemma61 _ _ _ _ H2).
+ apply lemma62.
+ split.
+ apply IHys.
+ intuition.
+ apply H0; intuition.
+
+ (* Z is complete *)
+ unfold complete.
+ unfold element_complete.
+ intros.
+ set (n:=enum x).
+ assert (Claim : equiconsistent (F_ n) (up (union_singl (F_ n) x))).
+ split.
+ unfold inconsistent.
+ intro.
+ apply lem223.
+ left.
+ assumption.
+ unfold inconsistent.
+ intro.
+ assert ((up (union_singl Z x)) bott).
+ apply subseteq_up_mono with (union_singl (F_ n) x).
+ intros.
+ unfold union_singl in H2 |- *.
+ intuition.
+ left.
+ exists n.
+ assumption.
+ assumption.
+ assert (Z bott).
+ clear -H0 H2.
+ firstorder.
+ assert (forall n : nat, equiconsistent (F_ n) Z).
+ clear -lem226 lem224.
+ intro n.
+ firstorder.
+ clear -H3 H4.
+ firstorder.
+ exists (S n).
+ simpl F_.
+ apply lem223.
+ right.
+ auto.
+ Defined.
+ End completion.
+
+ (** Some additional lemmas that are needed by classcomp.v *)
+ Section additional_lemmas.
+
+ Lemma lemma8 : forall (X:B -> Prop)(f:B)(ys:list B),
+ fold_left (fun (a : Prop) (b : B) => a /\ (union_singl X f) b) ys True ->
+ fold_left (fun (a : Prop) (b : B) => a /\ X b) ys True \/
+ fold_left (fun (a : Prop) (b : B) => a /\ X b)
+ (remove id_B_dec f ys) True
+ .
+ Proof.
+ fix 3.
+ destruct ys.
+ simpl.
+ intuition.
+ simpl.
+ intro.
+ apply lemma61 in H.
+ destruct H.
+ destruct (lemma8 _ _ _ H).
+ destruct H0.
+ left.
+ apply lemma62.
+ intuition.
+ right.
+ destruct (id_B_dec f b).
+ (* ys is a subset of ys\{f} *)
+ Focus 2.
+ congruence.
+ idtac.
+ Focus 2.
+ right.
+ destruct (id_B_dec f b).
+ assumption.
+ idtac.
+ simpl.
+ apply lemma62.
+ split.
+ assumption.
+ destruct H0.
+ assumption.
+ congruence.
+ idtac.
+ idtac.
+ clear -H1.
+ induction ys.
+ simpl in *.
+ auto.
+ simpl in *.
+ apply lemma61 in H1.
+ destruct (id_B_dec f a).
+ intuition.
+ simpl.
+ apply lemma62.
+ intuition.
+ Qed.
+
+ Lemma leq_bott : forall x:B, leq bott x.
+ Proof.
+ intro x.
+ unfold leq.
+ ring.
+ Qed.
+
+ End additional_lemmas.
+
+End filter_completion.
+
diff --git a/boolean_completeness/html/cPair.html b/boolean_completeness/html/cPair.html
new file mode 100644
index 0000000..d356cb5
--- /dev/null
+++ b/boolean_completeness/html/cPair.html
@@ -0,0 +1,133 @@
+
+
+
+
+
+cPair
+
+
+
+
+
+
+
+
+
+
+
Library cPair
+
+
+
+
+
+The Cantor pairing function proved correct.
+
+
+
+ This code is part of Russell O'Connor's formalisation of Goedel's
+ Incompleteness Theorem in Coq, released into the Public Domain.
+
+
+
+ http://r6.ca/Goedel/goedel1.html
+
+
+
+
Require Import Arith .
+
Require Import Coq.Lists.List .
+
Require Import extEqualNat .
+
Require Import primRec .
+
+
+
Definition sumToN (
n :
nat ) :=
+
nat_rec (
fun _ ⇒
nat ) 0 (
fun x y :
nat ⇒
S x + y )
n .
+
+
+
Lemma sumToN1 :
∀ n :
nat ,
n ≤ sumToN n .
+
+
+
Lemma sumToN2 :
∀ b a :
nat ,
a ≤ b →
sumToN a ≤ sumToN b .
+
+
+
Definition cPair (
a b :
nat ) :=
a + sumToN (
a + b ).
+
+
+
Section CPair_Injectivity .
+
+
+
Remark cPairInjHelp :
+
∀ a b c d :
nat ,
cPair a b = cPair c d →
a + b = c + d .
+
+
+
Lemma cPairInj1 :
∀ a b c d :
nat ,
cPair a b = cPair c d →
a = c .
+
+
+
Lemma cPairInj2 :
∀ a b c d :
nat ,
cPair a b = cPair c d →
b = d .
+
+
+
End CPair_Injectivity .
+
+
+
Section CPair_projections .
+
+
+
Let searchXY (
a :
nat ) :=
+
boundedSearch (
fun a y :
nat ⇒
ltBool a (
sumToN (
S y )))
a .
+
+
+
Definition cPairPi1 (
a :
nat ) :=
a - sumToN (
searchXY a ).
+
Definition cPairPi2 (
a :
nat ) :=
searchXY a - cPairPi1 a .
+
+
+
Lemma cPairProjectionsHelp :
+
∀ a b :
nat ,
b < sumToN (
S a ) →
sumToN a ≤ b →
searchXY b = a .
+
+
+
Lemma cPairProjections :
∀ a :
nat ,
cPair (
cPairPi1 a ) (
cPairPi2 a )
= a .
+
+
+
Lemma cPairProjections1 :
∀ a b :
nat ,
cPairPi1 (
cPair a b )
= a .
+
+
+
Lemma cPairProjections2 :
∀ a b :
nat ,
cPairPi2 (
cPair a b )
= b .
+
+
+
End CPair_projections .
+
+
+
Section CPair_Order .
+
+
+
Lemma cPairLe1 :
∀ a b :
nat ,
a ≤ cPair a b .
+
+
+
Lemma cPairLe1A :
∀ a :
nat ,
cPairPi1 a ≤ a .
+
+
+
Lemma cPairLe2 :
∀ a b :
nat ,
b ≤ cPair a b .
+
+
+
Lemma cPairLe2A :
∀ a :
nat ,
cPairPi2 a ≤ a .
+
+
+
Lemma cPairLe3 :
+
∀ a b c d :
nat ,
a ≤ b →
c ≤ d →
cPair a c ≤ cPair b d .
+
+
+
Lemma cPairLt1 :
∀ a b :
nat ,
a < cPair a (
S b ).
+
+
+
Lemma cPairLt2 :
∀ a b :
nat ,
b < cPair (
S a )
b .
+
+
+
End CPair_Order .
+
+
This page has been generated by
coqdoc
+
+
+
+
+
+
\ No newline at end of file
diff --git a/boolean_completeness/html/extEqualNat.html b/boolean_completeness/html/extEqualNat.html
new file mode 100644
index 0000000..e0b9691
--- /dev/null
+++ b/boolean_completeness/html/extEqualNat.html
@@ -0,0 +1,82 @@
+
+
+
+
+
+extEqualNat
+
+
+
+
+
+
+
+
+
+
+
Library extEqualNat
+
+
+
+
+
+This code is part of Russell O'Connor's formalisation of Goedel's
+ Incompleteness Theorem in Coq, released into the Public Domain.
+
+
+
+ http://r6.ca/Goedel/goedel1.html
+
+
+
+
This page has been generated by
coqdoc
+
+
+
+
+
+
\ No newline at end of file
diff --git a/boolean_completeness/html/index.html b/boolean_completeness/html/index.html
new file mode 100644
index 0000000..b4a86c1
--- /dev/null
+++ b/boolean_completeness/html/index.html
@@ -0,0 +1,22 @@
+
+
+Completeness for Boolean models formalised in Coq
+
+
+Completeness for Boolean models formalised in Coq
+
+The formal proofs are inside the archive boolean_completeness.zip . You can also browse the generated documentation as colored HTML code here .
+
+In order to compile the sources, you need the Coq proof assistant version 8.2.
+
+If you are on a Unix-like system, and the executable coqc is in your PATH variable, you can just run the command make from inside a terminal window. If coqc is not in your path, you can either set it, or open the Coq IDE and run Make from the Compile menu.
+
+If you are on a Windows system, you do not have make , and coqc is not in your PATH, the simplest solution is to open Coq IDE and compile one by one file, making sure to do it in the following order: pairing/extEqualNat.v, pairing/misc.v, pairing/primRec.v, pairing/cPair.v, lists.v, filters.v, classcomp.v.
+
+Note: The formalisation is not complete, sorting of lists in lists.v remains to be finished.
+
+
+Back to Danko Ilik's home page
+
+
+
diff --git a/boolean_completeness/html/misc.html b/boolean_completeness/html/misc.html
new file mode 100644
index 0000000..5c2f320
--- /dev/null
+++ b/boolean_completeness/html/misc.html
@@ -0,0 +1,49 @@
+
+
+
+
+
+misc
+
+
+
+
+
+
+
+
+
+
+
Library misc
+
+
+
+
+
+This code is part of Russell O'Connor's formalisation of Goedel's
+ Incompleteness Theorem in Coq, released into the Public Domain.
+
+
+
+ http://r6.ca/Goedel/goedel1.html
+
+
+
+
This page has been generated by
coqdoc
+
+
+
+
+
+
\ No newline at end of file
diff --git a/boolean_completeness/html/primRec.html b/boolean_completeness/html/primRec.html
new file mode 100644
index 0000000..cdc471d
--- /dev/null
+++ b/boolean_completeness/html/primRec.html
@@ -0,0 +1,713 @@
+
+
+
+
+
+primRec
+
+
+
+
+
+
+
+
+
+
+
Library primRec
+
+
+
+
+
+This code is part of Russell O'Connor's formalisation of Goedel's
+ Incompleteness Theorem in Coq, released into the Public Domain.
+
+
+
+ http://r6.ca/Goedel/goedel1.html
+
+
+
+
Require Import Arith .
+
Require Import Peano_dec .
+
Require Import Compare_dec .
+
Require Import Coq.Lists.List .
+
Require Import Eqdep_dec .
+
Require Import extEqualNat .
+
Require Import Bvector .
+
Require Import misc .
+
Require Export Bool .
+
Require Export EqNat .
+
Require Import Even .
+
Require Import Max .
+
+
+
Inductive PrimRec :
nat →
Set :=
+ |
succFunc :
PrimRec 1
+ |
zeroFunc :
PrimRec 0
+ |
projFunc :
∀ n m :
nat ,
m < n →
PrimRec n
+ |
composeFunc :
+
∀ (
n m :
nat ) (
g :
PrimRecs n m ) (
h :
PrimRec m ),
PrimRec n
+ |
primRecFunc :
+
∀ (
n :
nat ) (
g :
PrimRec n ) (
h :
PrimRec (
S (
S n ))),
PrimRec (
S n )
+
with PrimRecs :
nat →
nat →
Set :=
+ |
PRnil :
∀ n :
nat ,
PrimRecs n 0
+ |
PRcons :
∀ n m :
nat ,
PrimRec n →
PrimRecs n m →
PrimRecs n (
S m ).
+
+
+
Scheme PrimRec_PrimRecs_rec :=
Induction for PrimRec
+
Sort Set
+
with PrimRecs_PrimRec_rec :=
Induction for PrimRecs
+
Sort Set .
+
+
+
Scheme PrimRec_PrimRecs_ind :=
Induction for PrimRec
+
Sort Prop
+
with PrimRecs_PrimRec_ind :=
Induction for PrimRecs
+
Sort Prop .
+
+
+
Fixpoint evalConstFunc (
n m :
nat ) {
struct n } :
naryFunc n :=
+
match n return (
naryFunc n )
with
+ |
O ⇒
m
+ |
S n' ⇒
fun _ ⇒
evalConstFunc n' m
+
end .
+
+
+
Fixpoint evalProjFunc (
n :
nat ) :
∀ m :
nat ,
m < n →
naryFunc n :=
+
match n return (
∀ m :
nat ,
m < n →
naryFunc n )
with
+ |
O ⇒
fun (
m :
nat ) (
l :
m < 0) ⇒
False_rec _ (
lt_n_O _ l )
+ |
S n' ⇒
+
fun (
m :
nat ) (
l :
m < S n' ) ⇒
+
match eq_nat_dec m n' with
+ |
left _ ⇒
fun a :
nat ⇒
evalConstFunc _ a
+ |
right l1 ⇒
+
fun _ ⇒
+
evalProjFunc n' m
+
match le_lt_or_eq _ _ (
lt_n_Sm_le _ _ l )
with
+ |
or_introl l2 ⇒
l2
+ |
or_intror l2 ⇒
False_ind _ (
l1 l2 )
+
end
+
end
+
end .
+
+
+
Lemma evalProjFuncInd :
+
∀ (
n m :
nat ) (
p1 p2 :
m < n ),
+
evalProjFunc n m p1 = evalProjFunc n m p2 .
+
+
+
Fixpoint evalList (
m :
nat ) (
l :
Vector.t nat m ) {
struct l } :
+
naryFunc m →
nat :=
+
match l with
+ |
Vector.nil ⇒
fun x :
naryFunc 0 ⇒
x
+ |
Vector.cons a n l' ⇒
fun x :
naryFunc (
S n ) ⇒
evalList n l' (
x a )
+
end .
+
+
+
Fixpoint evalOneParamList (
n m a :
nat ) (
l :
Vector.t (
naryFunc (
S n ))
m )
+ {
struct l } :
Vector.t (
naryFunc n )
m :=
+
match l in (
Vector.t _ m )
return (
Vector.t (
naryFunc n )
m )
with
+ |
Vector.nil ⇒
Vector.nil (
naryFunc n )
+ |
Vector.cons f m' l' ⇒
Vector.cons _ (
f a )
m' (
evalOneParamList n m' a l' )
+
end .
+
+
+
Fixpoint evalComposeFunc (
n :
nat ) :
+
∀ m :
nat ,
Vector.t (
naryFunc n )
m →
naryFunc m →
naryFunc n :=
+
match
+
n
+
return
+ (
∀ m :
nat ,
Vector.t (
naryFunc n )
m →
naryFunc m →
naryFunc n )
+
with
+ |
O ⇒
evalList
+ |
S n' ⇒
+
fun (
m :
nat ) (
l :
Vector.t (
naryFunc (
S n' ))
m )
+ (
f :
naryFunc m ) (
a :
nat ) ⇒
+
evalComposeFunc n' m (
evalOneParamList _ _ a l )
f
+
end .
+
+
+
Fixpoint compose2 (
n :
nat ) :
naryFunc n →
naryFunc (
S n ) →
naryFunc n :=
+
match n return (
naryFunc n →
naryFunc (
S n ) →
naryFunc n )
with
+ |
O ⇒
fun (
a :
nat ) (
g :
nat →
nat ) ⇒
g a
+ |
S n' ⇒
+
fun (
f :
naryFunc (
S n' )) (
g :
naryFunc (
S (
S n' ))) (
a :
nat ) ⇒
+
compose2 n' (
f a ) (
fun x :
nat ⇒
g x a )
+
end .
+
+
+
Fixpoint evalPrimRecFunc (
n :
nat ) (
g :
naryFunc n )
+ (
h :
naryFunc (
S (
S n ))) (
a :
nat ) {
struct a } :
naryFunc n :=
+
match a with
+ |
O ⇒
g
+ |
S a' ⇒
compose2 _ (
evalPrimRecFunc n g h a' ) (
h a' )
+
end .
+
+
+
Fixpoint evalPrimRec (
n :
nat ) (
f :
PrimRec n ) {
struct f } :
+
naryFunc n :=
+
match f in (
PrimRec n )
return (
naryFunc n )
with
+ |
succFunc ⇒
S
+ |
zeroFunc ⇒ 0
+ |
projFunc n m pf ⇒
evalProjFunc n m pf
+ |
composeFunc n m l f ⇒
+
evalComposeFunc n m (
evalPrimRecs _ _ l ) (
evalPrimRec _ f )
+ |
primRecFunc n g h ⇒
+
evalPrimRecFunc n (
evalPrimRec _ g ) (
evalPrimRec _ h )
+
end
+
+
with evalPrimRecs (
n m :
nat ) (
fs :
PrimRecs n m ) {
struct fs } :
+
Vector.t (
naryFunc n )
m :=
+
match fs in (
PrimRecs n m )
return (
Vector.t (
naryFunc n )
m )
with
+ |
PRnil a ⇒
Vector.nil (
naryFunc a )
+ |
PRcons a b g gs ⇒
Vector.cons _ (
evalPrimRec _ g )
_ (
evalPrimRecs _ _ gs )
+
end .
+
+
+
Definition extEqualVectorGeneral :
∀ (
n m :
nat ) (
l :
Vector.t (
naryFunc n )
m )
+ (
m' :
nat ) (
l' :
Vector.t (
naryFunc n )
m' ) ,
Prop .
+
Defined .
+
+
+
Definition extEqualVector :
∀ (
n m :
nat ) (
l l' :
Vector.t (
naryFunc n )
m ) ,
Prop .
+
Defined .
+
+
+
Lemma extEqualVectorRefl :
+
∀ (
n m :
nat ) (
l :
Vector.t (
naryFunc n )
m ),
extEqualVector n m l l .
+
+
+
Lemma extEqualOneParamList :
+
∀ (
n m :
nat ) (
l1 l2 :
Vector.t (
naryFunc (
S n ))
m ) (
c :
nat ),
+
extEqualVector (
S n )
m l1 l2 →
+
extEqualVector n m (
evalOneParamList n m c l1 ) (
evalOneParamList n m c l2 ).
+
+
+
Lemma extEqualCompose :
+
∀ (
n m :
nat ) (
l1 l2 :
Vector.t (
naryFunc n )
m ) (
f1 f2 :
naryFunc m ),
+
extEqualVector n m l1 l2 →
+
extEqual m f1 f2 →
+
extEqual n (
evalComposeFunc n m l1 f1 ) (
evalComposeFunc n m l2 f2 ).
+
Definition Vhead {
T :
Type } (
n :
nat ) (
v :
Vector.t T (
S n )) :=
+
match v with
+ |
Vector.cons a _ _ ⇒
a
+
end .
+
+
+
Definition Vtail {
T :
Type } (
n :
nat ) (
v :
Vector.t T (
S n )) :=
+
match v with
+ |
Vector.cons _ _ v ⇒
v
+
end .
+
+
+
Lemma V0_eq :
∀ (
T :
Type ) (
v :
Vector.t T 0),
v = Vector.nil T .
+
Print Vector.t .
+
Qed .
+
+
+
Lemma VSn_eq :
+
∀ (
T :
Type ) (
n :
nat ) (
v :
Vector.t T (
S n )),
v = Vector.cons T (
Vhead _ v )
_ (
Vtail _ v ).
+
+
+
Qed .
+
+
+
Lemma extEqualCompose2 :
+
∀ (
n :
nat ) (
f1 f2 :
naryFunc n ),
+
extEqual n f1 f2 →
+
∀ g1 g2 :
naryFunc (
S n ),
+
extEqual (
S n )
g1 g2 →
extEqual n (
compose2 n f1 g1 ) (
compose2 n f2 g2 ).
+
+
+
Lemma extEqualPrimRec :
+
∀ (
n :
nat ) (
g1 g2 :
naryFunc n ) (
h1 h2 :
naryFunc (
S (
S n ))),
+
extEqual n g1 g2 →
+
extEqual (
S (
S n ))
h1 h2 →
+
extEqual (
S n ) (
evalPrimRecFunc n g1 h1 ) (
evalPrimRecFunc n g2 h2 ).
+
+
+
Definition isPR (
n :
nat ) (
f :
naryFunc n ) :
Set :=
+
{ p : PrimRec n | extEqual n (
evalPrimRec _ p )
f } .
+
+
+
Definition isPRrel (
n :
nat ) (
R :
naryRel n ) :
Set :=
+
isPR n (
charFunction n R ).
+
+
+
Lemma succIsPR :
isPR 1
S .
+
+
+
Lemma const0_NIsPR :
∀ n :
nat ,
isPR 0
n .
+
+
+
Lemma const1_NIsPR :
∀ n :
nat ,
isPR 1 (
fun _ ⇒
n ).
+
+
+
Lemma idIsPR :
isPR 1 (
fun x :
nat ⇒
x ).
+
+
+
Lemma pi1_2IsPR :
isPR 2 (
fun a b :
nat ⇒
a ).
+
+
+
Lemma pi2_2IsPR :
isPR 2 (
fun a b :
nat ⇒
b ).
+
+
+
Lemma pi1_3IsPR :
isPR 3 (
fun a b c :
nat ⇒
a ).
+
+
+
Lemma pi2_3IsPR :
isPR 3 (
fun a b c :
nat ⇒
b ).
+
+
+
Lemma pi3_3IsPR :
isPR 3 (
fun a b c :
nat ⇒
c ).
+
+
+
Lemma pi1_4IsPR :
isPR 4 (
fun a b c d :
nat ⇒
a ).
+
+
+
Lemma pi2_4IsPR :
isPR 4 (
fun a b c d :
nat ⇒
b ).
+
+
+
Lemma pi3_4IsPR :
isPR 4 (
fun a b c d :
nat ⇒
c ).
+
+
+
Lemma pi4_4IsPR :
isPR 4 (
fun a b c d :
nat ⇒
d ).
+
+
+
Lemma filter01IsPR :
+
∀ g :
nat →
nat ,
isPR 1
g →
isPR 2 (
fun a b :
nat ⇒
g b ).
+
+
+
Lemma filter10IsPR :
+
∀ g :
nat →
nat ,
isPR 1
g →
isPR 2 (
fun a b :
nat ⇒
g a ).
+
+
+
Lemma filter100IsPR :
+
∀ g :
nat →
nat ,
isPR 1
g →
isPR 3 (
fun a b c :
nat ⇒
g a ).
+
+
+
Lemma filter010IsPR :
+
∀ g :
nat →
nat ,
isPR 1
g →
isPR 3 (
fun a b c :
nat ⇒
g b ).
+
+
+
Lemma filter001IsPR :
+
∀ g :
nat →
nat ,
isPR 1
g →
isPR 3 (
fun a b c :
nat ⇒
g c ).
+
+
+
Lemma filter011IsPR :
+
∀ g :
nat →
nat →
nat ,
isPR 2
g →
isPR 3 (
fun a b c :
nat ⇒
g b c ).
+
+
+
Lemma filter110IsPR :
+
∀ g :
nat →
nat →
nat ,
isPR 2
g →
isPR 3 (
fun a b c :
nat ⇒
g a b ).
+
+
+
Lemma filter101IsPR :
+
∀ g :
nat →
nat →
nat ,
isPR 2
g →
isPR 3 (
fun a b c :
nat ⇒
g a c ).
+
+
+
Lemma filter0011IsPR :
+
∀ g :
nat →
nat →
nat ,
+
isPR 2
g →
isPR 4 (
fun a b c d :
nat ⇒
g c d ).
+
+
+
Lemma filter1000IsPR :
+
∀ g :
nat →
nat ,
isPR 1
g →
isPR 4 (
fun a b c d :
nat ⇒
g a ).
+
+
+
Lemma filter1011IsPR :
+
∀ g :
nat →
nat →
nat →
nat ,
+
isPR 3
g →
isPR 4 (
fun a b c d :
nat ⇒
g a c d ).
+
+
+
Lemma filter1100IsPR :
+
∀ g :
nat →
nat →
nat ,
+
isPR 2
g →
isPR 4 (
fun a b c d :
nat ⇒
g a b ).
+
+
+
Lemma compose1_1IsPR :
+
∀ f :
nat →
nat ,
+
isPR 1
f →
+
∀ g :
nat →
nat ,
isPR 1
g →
isPR 1 (
fun x :
nat ⇒
g (
f x )).
+
+
+
Lemma compose1_2IsPR :
+
∀ f :
nat →
nat ,
+
isPR 1
f →
+
∀ f' :
nat →
nat ,
+
isPR 1
f' →
+
∀ g :
nat →
nat →
nat ,
+
isPR 2
g →
isPR 1 (
fun x :
nat ⇒
g (
f x ) (
f' x )).
+
+
+
Lemma compose1_3IsPR :
+
∀ f1 :
nat →
nat ,
+
isPR 1
f1 →
+
∀ f2 :
nat →
nat ,
+
isPR 1
f2 →
+
∀ f3 :
nat →
nat ,
+
isPR 1
f3 →
+
∀ g :
nat →
nat →
nat →
nat ,
+
isPR 3
g →
isPR 1 (
fun x :
nat ⇒
g (
f1 x ) (
f2 x ) (
f3 x )).
+
+
+
Lemma compose2_1IsPR :
+
∀ f :
nat →
nat →
nat ,
+
isPR 2
f →
+
∀ g :
nat →
nat ,
isPR 1
g →
isPR 2 (
fun x y :
nat ⇒
g (
f x y )).
+
+
+
Lemma compose2_2IsPR :
+
∀ f :
nat →
nat →
nat ,
+
isPR 2
f →
+
∀ g :
nat →
nat →
nat ,
+
isPR 2
g →
+
∀ h :
nat →
nat →
nat ,
+
isPR 2
h →
isPR 2 (
fun x y :
nat ⇒
h (
f x y ) (
g x y )).
+
+
+
Lemma compose2_3IsPR :
+
∀ f1 :
nat →
nat →
nat ,
+
isPR 2
f1 →
+
∀ f2 :
nat →
nat →
nat ,
+
isPR 2
f2 →
+
∀ f3 :
nat →
nat →
nat ,
+
isPR 2
f3 →
+
∀ g :
nat →
nat →
nat →
nat ,
+
isPR 3
g →
isPR 2 (
fun x y :
nat ⇒
g (
f1 x y ) (
f2 x y ) (
f3 x y )).
+
+
+
Lemma compose2_4IsPR :
+
∀ f1 :
nat →
nat →
nat ,
+
isPR 2
f1 →
+
∀ f2 :
nat →
nat →
nat ,
+
isPR 2
f2 →
+
∀ f3 :
nat →
nat →
nat ,
+
isPR 2
f3 →
+
∀ f4 :
nat →
nat →
nat ,
+
isPR 2
f4 →
+
∀ g :
nat →
nat →
nat →
nat →
nat ,
+
isPR 4
g →
isPR 2 (
fun x y :
nat ⇒
g (
f1 x y ) (
f2 x y ) (
f3 x y ) (
f4 x y )).
+
+
+
Lemma compose3_1IsPR :
+
∀ f :
nat →
nat →
nat →
nat ,
+
isPR 3
f →
+
∀ g :
nat →
nat ,
isPR 1
g →
isPR 3 (
fun x y z :
nat ⇒
g (
f x y z )).
+
+
+
Lemma compose3_2IsPR :
+
∀ f1 :
nat →
nat →
nat →
nat ,
+
isPR 3
f1 →
+
∀ f2 :
nat →
nat →
nat →
nat ,
+
isPR 3
f2 →
+
∀ g :
nat →
nat →
nat ,
+
isPR 2
g →
isPR 3 (
fun x y z :
nat ⇒
g (
f1 x y z ) (
f2 x y z )).
+
+
+
Lemma compose3_3IsPR :
+
∀ f1 :
nat →
nat →
nat →
nat ,
+
isPR 3
f1 →
+
∀ f2 :
nat →
nat →
nat →
nat ,
+
isPR 3
f2 →
+
∀ f3 :
nat →
nat →
nat →
nat ,
+
isPR 3
f3 →
+
∀ g :
nat →
nat →
nat →
nat ,
+
isPR 3
g →
isPR 3 (
fun x y z :
nat ⇒
g (
f1 x y z ) (
f2 x y z ) (
f3 x y z )).
+
+
+
Lemma compose4_2IsPR :
+
∀ f1 :
nat →
nat →
nat →
nat →
nat ,
+
isPR 4
f1 →
+
∀ f2 :
nat →
nat →
nat →
nat →
nat ,
+
isPR 4
f2 →
+
∀ g :
nat →
nat →
nat ,
+
isPR 2
g →
isPR 4 (
fun w x y z :
nat ⇒
g (
f1 w x y z ) (
f2 w x y z )).
+
+
+
Lemma compose4_3IsPR :
+
∀ f1 :
nat →
nat →
nat →
nat →
nat ,
+
isPR 4
f1 →
+
∀ f2 :
nat →
nat →
nat →
nat →
nat ,
+
isPR 4
f2 →
+
∀ f3 :
nat →
nat →
nat →
nat →
nat ,
+
isPR 4
f3 →
+
∀ g :
nat →
nat →
nat →
nat ,
+
isPR 3
g →
+
isPR 4 (
fun w x y z :
nat ⇒
g (
f1 w x y z ) (
f2 w x y z ) (
f3 w x y z )).
+
+
+
Lemma swapIsPR :
+
∀ f :
nat →
nat →
nat ,
isPR 2
f →
isPR 2 (
fun x y :
nat ⇒
f y x ).
+
+
+
Lemma indIsPR :
+
∀ f :
nat →
nat →
nat ,
+
isPR 2
f →
+
∀ g :
nat ,
+
isPR 1
+ (
fun a :
nat ⇒
nat_rec (
fun n :
nat ⇒
nat )
g (
fun x y :
nat ⇒
f x y )
a ).
+
+
+
Lemma ind1ParamIsPR :
+
∀ f :
nat →
nat →
nat →
nat ,
+
isPR 3
f →
+
∀ g :
nat →
nat ,
+
isPR 1
g →
+
isPR 2
+ (
fun a b :
nat ⇒
+
nat_rec (
fun n :
nat ⇒
nat ) (
g b ) (
fun x y :
nat ⇒
f x y b )
a ).
+
+
+
Lemma ind2ParamIsPR :
+
∀ f :
nat →
nat →
nat →
nat →
nat ,
+
isPR 4
f →
+
∀ g :
nat →
nat →
nat ,
+
isPR 2
g →
+
isPR 3
+ (
fun a b c :
nat ⇒
+
nat_rec (
fun n :
nat ⇒
nat ) (
g b c ) (
fun x y :
nat ⇒
f x y b c )
a ).
+
+
+
Lemma plusIndIsPR :
isPR 3 (
fun n fn b :
nat ⇒
S fn ).
+
+
+
Lemma plusIsPR :
isPR 2
plus .
+
+
+
Lemma multIndIsPR :
isPR 3 (
fun n fn b :
nat ⇒
fn + b ).
+
+
+
Lemma multIsPR :
isPR 2
mult .
+
+
+
Lemma predIsPR :
isPR 1
pred .
+
+
+
Lemma minusIndIsPR :
isPR 3 (
fun n fn b :
nat ⇒
pred fn ).
+
+
+
Lemma minusIsPR :
isPR 2
minus .
+
+
+
Definition notZero (
a :
nat ) :=
+
nat_rec (
fun n :
nat ⇒
nat ) 0 (
fun x y :
nat ⇒ 1)
a .
+
+
+
Lemma notZeroIsPR :
isPR 1
notZero .
+
+
+
Definition ltBool (
a b :
nat ) :
bool .
+
Defined .
+
+
+
Lemma ltBoolTrue :
∀ a b :
nat ,
ltBool a b = true →
a < b .
+
+
+
Lemma ltBoolFalse :
∀ a b :
nat ,
ltBool a b = false →
¬ a < b .
+
+
+
Lemma ltIsPR :
isPRrel 2
ltBool .
+
+
+
Lemma maxIsPR :
isPR 2
max .
+
+
+
Lemma gtIsPR :
isPRrel 2 (
fun a b :
nat ⇒
ltBool b a ).
+
+
+
Remark replaceCompose2 :
+
∀ (
n :
nat ) (
a b a' b' :
naryFunc n ) (
c c' :
naryFunc 2),
+
extEqual n a a' →
+
extEqual n b b' →
+
extEqual 2
c c' →
+
extEqual n
+ (
evalComposeFunc _ _ (
Vector.cons _ a _ (
Vector.cons _ b _ (
Vector.nil (
naryFunc n ))))
c )
+ (
evalComposeFunc _ _ (
Vector.cons _ a' _ (
Vector.cons _ b' _ (
Vector.nil (
naryFunc n ))))
c' ).
+
+
+
Definition orRel (
n :
nat ) (
R S :
naryRel n ) :
naryRel n .
+
Defined .
+
+
+
Lemma orRelPR :
+
∀ (
n :
nat ) (
R R' :
naryRel n ),
+
isPRrel n R →
isPRrel n R' →
isPRrel n (
orRel n R R' ).
+
+
+
Definition andRel (
n :
nat ) (
R S :
naryRel n ) :
naryRel n .
+
Defined .
+
+
+
Lemma andRelPR :
+
∀ (
n :
nat ) (
R R' :
naryRel n ),
+
isPRrel n R →
isPRrel n R' →
isPRrel n (
andRel n R R' ).
+
+
+
Definition notRel (
n :
nat ) (
R :
naryRel n ) :
naryRel n .
+
Defined .
+
+
+
Lemma notRelPR :
+
∀ (
n :
nat ) (
R :
naryRel n ),
isPRrel n R →
isPRrel n (
notRel n R ).
+
+
+
Fixpoint bodd (
n :
nat ) :
bool :=
+
match n with
+ |
O ⇒
false
+ |
S n' ⇒
negb (
bodd n' )
+
end .
+
+
+
Lemma boddIsPR :
isPRrel 1
bodd .
+
+
+
Lemma beq_nat_not_refl :
∀ a b :
nat ,
a ≠ b →
beq_nat a b = false .
+
+
+
Lemma neqIsPR :
isPRrel 2 (
fun a b :
nat ⇒
negb (
beq_nat a b )).
+
+
+
Lemma eqIsPR :
isPRrel 2
beq_nat .
+
+
+
Definition leBool (
a b :
nat ) :
bool .
+
Defined .
+
+
+
Lemma leIsPR :
isPRrel 2
leBool .
+
+
+
Section Ignore_Params .
+
+
+
Fixpoint ignoreParams (
n m :
nat ) (
f :
naryFunc n ) {
struct m } :
+
naryFunc (
m + n ) :=
+
match m return (
naryFunc (
m + n ))
with
+ |
O ⇒
f
+ |
S x ⇒
fun _ ⇒
ignoreParams n x f
+
end .
+
+
+
Definition projectionListPR (
n m :
nat ) (
p :
m ≤ n ) :
PrimRecs n m .
+
Defined .
+
+
+
Definition projectionList (
n m :
nat ) (
p :
m ≤ n ) :
+
Vector.t (
naryFunc n )
m :=
evalPrimRecs n m (
projectionListPR n m p ).
+
+
+
Lemma projectionListInd :
+
∀ (
n m :
nat ) (
p1 p2 :
m ≤ n ),
+
projectionList n m p1 = projectionList n m p2 .
+
+
+
Lemma projectionListApplyParam :
+
∀ (
m n c :
nat ) (
p1 :
m ≤ n ) (
p2 :
m ≤ S n ),
+
extEqualVector _ _ (
projectionList n m p1 )
+ (
evalOneParamList n m c (
projectionList (
S n )
m p2 )).
+
+
+
Lemma projectionListId :
+
∀ (
n :
nat ) (
f :
naryFunc n ) (
p :
n ≤ n ),
+
extEqual n f (
evalComposeFunc n n (
projectionList n n p )
f ).
+
+
+
Lemma ignoreParamsIsPR :
+
∀ (
n m :
nat ) (
f :
naryFunc n ),
isPR _ f →
isPR _ (
ignoreParams n m f ).
+
+
+
End Ignore_Params .
+
+
+
Lemma reduce1stCompose :
+
∀ (
c n m :
nat ) (
v :
Vector.t (
naryFunc n )
m ) (
g :
naryFunc (
S m )),
+
extEqual n
+ (
evalComposeFunc n _ (
Vector.cons (
naryFunc n ) (
evalConstFunc n c )
_ v )
g )
+ (
evalComposeFunc n _ v (
g c )).
+
+
+
Lemma reduce2ndCompose :
+
∀ (
c n m :
nat ) (
v :
Vector.t (
naryFunc n )
m ) (
n0 :
naryFunc n )
+ (
g :
naryFunc (
S (
S m ))),
+
extEqual n
+ (
evalComposeFunc n _
+ (
Vector.cons (
naryFunc n )
n0 _ (
Vector.cons (
naryFunc n ) (
evalConstFunc n c )
_ v ))
+
g )
+ (
evalComposeFunc n _ (
Vector.cons (
naryFunc n )
n0 _ v ) (
fun x :
nat ⇒
g x c )).
+
+
+
Lemma evalPrimRecParam :
+
∀ (
n c d :
nat ) (
g :
naryFunc (
S n )) (
h :
naryFunc (
S (
S (
S n )))),
+
extEqual _ (
evalPrimRecFunc n (
g d ) (
fun x y :
nat ⇒
h x y d )
c )
+ (
evalPrimRecFunc (
S n )
g h c d ).
+
+
+
Lemma compose2IsPR :
+
∀ (
n :
nat ) (
f :
naryFunc n ) (
p :
isPR n f ) (
g :
naryFunc (
S n ))
+ (
q :
isPR (
S n )
g ),
isPR n (
compose2 n f g ).
+
+
+
Lemma compose1_NIsPR :
+
∀ (
n :
nat ) (
g :
naryFunc (
S n )),
+
isPR (
S n )
g →
+
∀ f :
naryFunc 1,
isPR 1
f →
isPR (
S n ) (
fun x :
nat ⇒
g (
f x )).
+
+
+
Definition switchPR :
naryFunc 3.
+
Defined .
+
+
+
Lemma switchIsPR :
isPR 3
switchPR .
+
+
+
Fixpoint boundedSearchHelp (
P :
naryRel 1) (
b :
nat ) {
struct b } :
nat :=
+
match b with
+ |
O ⇒ 0
+ |
S b' ⇒
+
match eq_nat_dec (
boundedSearchHelp P b' )
b' with
+ |
left _ ⇒
match P b' with
+ |
true ⇒
b'
+ |
false ⇒
S b'
+
end
+ |
right _ ⇒
boundedSearchHelp P b'
+
end
+
end .
+
+
+
Definition boundedSearch (
P :
naryRel 2) (
b :
nat ) :
nat :=
+
boundedSearchHelp (
P b )
b .
+
+
+
Lemma boundedSearch1 :
+
∀ (
P :
naryRel 2) (
b x :
nat ),
x < boundedSearch P b →
P b x = false .
+
+
+
Lemma boundedSearch2 :
+
∀ (
P :
naryRel 2) (
b :
nat ),
+
boundedSearch P b = b ∨ P b (
boundedSearch P b )
= true .
+
+
+
Lemma boundSearchIsPR :
+
∀ P :
naryRel 2,
+
isPRrel 2
P →
isPR 1 (
fun b :
nat ⇒
boundedSearch P b ).
+
+
+
Definition iterate (
g :
nat →
nat ) :=
+
nat_rec (
fun _ ⇒
nat →
nat ) (
fun x :
nat ⇒
x )
+ (
fun (
_ :
nat ) (
rec :
nat →
nat ) (
x :
nat ) ⇒
g (
rec x )).
+
+
+
Lemma iterateIsPR :
+
∀ g :
nat →
nat ,
isPR 1
g →
∀ n :
nat ,
isPR 1 (
iterate g n ).
+
+
This page has been generated by
coqdoc
+
+
+
+
+
+
\ No newline at end of file
diff --git a/boolean_completeness/lists.v b/boolean_completeness/lists.v
new file mode 100644
index 0000000..8bce38c
--- /dev/null
+++ b/boolean_completeness/lists.v
@@ -0,0 +1,225 @@
+(** A file that involves general list sorting with respect to a
+ nonstandard specification. The proofs are not finished. *)
+
+Require Import List.
+
+Set Implicit Arguments.
+
+Section ListExtras.
+Variable T:Type.
+
+(* this is essentially lelist from the stdlib, but there are not many
+ results proved about it there, so I don't know if it's worth
+ switching *)
+Inductive Suffix : list T -> list T -> Prop :=
+| Suffix_refl: forall xs, Suffix xs xs
+| Suffix_cons: forall x xs ys, Suffix ys xs -> Suffix ys (x::xs).
+
+Lemma Suffix_lem1 : forall xs ys:list T, Suffix xs ys ->
+ exists zs, zs++xs=ys.
+Proof.
+ intros.
+ induction ys.
+ inversion H.
+ exists nil.
+ simpl.
+ reflexivity.
+ inversion H.
+ exists nil.
+ simpl.
+ reflexivity.
+ assert (IH := IHys H2).
+ clear -IH.
+ destruct IH as [l Hl].
+ rewrite <- Hl.
+ exists (a::l).
+ simpl.
+ reflexivity.
+Qed.
+
+Lemma Suffix_incl : forall xs ys:list T, Suffix xs ys ->
+ incl xs ys.
+Proof.
+ intros.
+ induction ys.
+ inversion H.
+ red.
+ auto.
+ inversion H.
+ red.
+ auto.
+ red.
+ assert (IH:=IHys H2).
+ clear -IH.
+ red in IH.
+ auto using in_cons.
+Qed.
+
+Section Sorting.
+
+ Variable R:T->T->bool.
+
+ (* sorting is taken from the CoLoR library *)
+ Fixpoint insert x (l : list T) :=
+ match l with
+ | nil => cons x nil
+ | cons y m =>
+ if R x y then (cons y (insert x m)) else (cons x l)
+ end.
+
+ Fixpoint sort (l : list T) : list T :=
+ match l with
+ | nil => nil
+ | cons x m => insert x (sort m)
+ end.
+
+ Lemma insert_ext : forall xs a a0,
+ a = a0 \/ In a0 xs <-> In a0 (insert a xs).
+ Proof.
+ induction xs.
+ simpl.
+ intuition.
+
+ simpl insert.
+ intros b c.
+ Require Setoid.
+ destruct R.
+ simpl.
+ rewrite <- IHxs.
+ intuition.
+ simpl.
+ intuition.
+ Qed.
+
+ Lemma insert_sort_ext : forall xs a a0,
+ a = a0 \/ In a0 xs <-> In a0 (insert a (sort xs)).
+ Proof.
+ induction xs.
+ simpl.
+ intuition.
+
+ simpl insert.
+ intros b c.
+ Require Setoid.
+ simpl.
+ rewrite <- insert_ext.
+ rewrite <- IHxs.
+ intuition.
+ Qed.
+
+ Lemma sort_ext : forall (xs:list T),
+ let ys := sort xs in
+ incl xs ys /\ incl ys xs.
+ Proof.
+ intros xs ys.
+
+ split.
+ induction xs.
+ intuition.
+ unfold ys in IHxs.
+ unfold ys.
+ clear -IHxs.
+ simpl.
+ unfold incl in *.
+ simpl in *.
+ intros.
+ firstorder using insert_sort_ext.
+
+ induction xs.
+ intuition.
+ simpl in IHxs.
+ unfold ys.
+ clear -IHxs.
+ simpl.
+ unfold incl in *.
+ simpl in *.
+ intros.
+ firstorder using insert_sort_ext.
+ Qed.
+
+ Lemma sort_suffix :
+ forall xs zs z, Suffix (z::zs) (sort xs) ->
+ forall z', In z' zs -> R z z' = true.
+ Proof.
+ intros xs zs z H.
+ assert ({ys:list T|ys=z::zs}).
+ exists (z::zs).
+ reflexivity.
+ destruct X as [ys Hys].
+ rewrite <-Hys in H.
+ replace zs with (tail ys).
+ Focus 2.
+ rewrite Hys; simpl; reflexivity.
+ apply Suffix_lem1 in H.
+ generalize dependent xs.
+ generalize dependent zs.
+ induction ys.
+ intuition.
+ simpl.
+ intros.
+ assert (H1 : ys=nil \/ In z' (tail ys) \/ Some z' = head ys).
+ admit.
+ destruct H1.
+ admit.
+ destruct H1.
+ Admitted.
+
+ Lemma sort_correct : forall (xs:list T),
+ let ys := sort xs in
+ incl xs ys /\ incl ys xs /\
+ forall z zs, Suffix (z::zs) ys ->
+ forall z', In z' zs -> R z z' = true.
+ Proof.
+ intros xs ys.
+
+ split.
+ firstorder using sort_ext.
+ split.
+ firstorder using sort_ext.
+ eauto using sort_suffix.
+ Qed.
+
+ Lemma nodup_sort : forall (xs:list T),
+ NoDup xs -> NoDup (sort xs).
+ Admitted.
+
+
+(* I should have a look at the place where this list operations are
+ used. Maybe it's possible to optimise a bit. Ex: do we compose a
+ nodup with sort or vice-versa. *)
+
+End Sorting.
+
+Section Removing_consecutive_duplicates.
+
+ Variable R:T->T->bool.
+ Variable eqT_dec:forall x y:T, {x=y}+{x<>y}.
+
+ (* remove _consecutive_ appearences of duplicate elements of a list *)
+ Fixpoint nodup' (xs:list T) : list T :=
+ match xs with
+ | nil => nil
+ | cons y ys => cons y
+ match ys with
+ | nil => nil
+ | cons z zs =>
+ if eqT_dec y z then (nodup' zs) else cons z (nodup' zs)
+ end
+ end.
+
+ Definition nodup (xs:list T) : list T := nodup' (sort R xs).
+
+ Lemma nodup_correct : forall xs:list T,
+ let ys := nodup xs in
+ incl xs ys /\ incl ys xs /\ NoDup ys.
+ Admitted.
+
+End Removing_consecutive_duplicates.
+
+
+End ListExtras.
+
+(* Require Import Compare_dec. *)
+(* Require Import Peano_dec. *)
+(* Eval compute in (sort leb (1::2::1::4::3::4::nil)). *)
+(* Eval compute in (nodup leb eq_nat_dec (1::2::1::4::3::4::nil)). *)
diff --git a/boolean_completeness/pairing/cPair.v b/boolean_completeness/pairing/cPair.v
new file mode 100644
index 0000000..89ff24b
--- /dev/null
+++ b/boolean_completeness/pairing/cPair.v
@@ -0,0 +1,339 @@
+(** The Cantor pairing function proved correct.
+
+ This code is part of Russell O'Connor's formalisation of Goedel's
+ Incompleteness Theorem in Coq, released into the Public Domain.
+
+ http://r6.ca/Goedel/goedel1.html
+*)
+Require Import Arith.
+Require Import Coq.Lists.List.
+Require Import extEqualNat.
+Require Import primRec.
+
+Definition sumToN (n : nat) :=
+ nat_rec (fun _ => nat) 0 (fun x y : nat => S x + y) n.
+
+Lemma sumToN1 : forall n : nat, n <= sumToN n.
+Proof.
+intros.
+induction n as [| n Hrecn].
+auto.
+simpl in |- *.
+apply le_n_S.
+apply le_plus_l.
+Qed.
+
+Lemma sumToN2 : forall b a : nat, a <= b -> sumToN a <= sumToN b.
+Proof.
+intro.
+induction b as [| b Hrecb]; intros.
+simpl in |- *.
+rewrite <- (le_n_O_eq _ H).
+simpl in |- *.
+auto.
+induction (le_lt_or_eq _ _ H).
+apply le_trans with (sumToN b).
+apply Hrecb.
+apply lt_n_Sm_le.
+auto.
+simpl in |- *.
+apply le_S.
+apply le_plus_r.
+rewrite H0.
+auto.
+Qed.
+
+Definition cPair (a b : nat) := a + sumToN (a + b).
+
+Section CPair_Injectivity.
+
+Remark cPairInjHelp :
+ forall a b c d : nat, cPair a b = cPair c d -> a + b = c + d.
+Proof.
+assert (forall a b : nat, a < b -> a + sumToN a < sumToN b).
+simple induction b.
+intros.
+elim (lt_n_O _ H).
+intros.
+simpl in |- *.
+assert (a <= n).
+apply lt_n_Sm_le.
+assumption.
+induction (le_lt_or_eq a n H1).
+apply lt_trans with (sumToN n).
+auto.
+apply le_lt_n_Sm.
+apply le_plus_r.
+rewrite H2.
+apply lt_n_Sn.
+unfold cPair in |- *.
+assert
+ (forall a b c d : nat,
+ a <= c -> b <= d -> a + sumToN c = b + sumToN d -> c = d).
+intros.
+induction (le_or_lt c d).
+induction (le_lt_or_eq _ _ H3).
+assert (a + sumToN c < sumToN d).
+apply le_lt_trans with (c + sumToN c).
+apply plus_le_compat_r.
+auto.
+auto.
+rewrite H2 in H5.
+elim (lt_not_le _ _ H5).
+apply le_plus_r.
+auto.
+assert (b + sumToN d < sumToN c).
+apply le_lt_trans with (d + sumToN d).
+apply plus_le_compat_r.
+auto.
+auto.
+rewrite <- H2 in H4.
+elim (lt_not_le _ _ H4).
+apply le_plus_r.
+intros.
+eapply H0.
+apply le_plus_l.
+apply le_plus_l.
+auto.
+Qed.
+
+Lemma cPairInj1 : forall a b c d : nat, cPair a b = cPair c d -> a = c.
+Proof.
+intros.
+assert (a + b = c + d).
+apply cPairInjHelp.
+auto.
+eapply plus_reg_l.
+unfold cPair in H.
+rewrite (plus_comm a) in H.
+rewrite (plus_comm c) in H.
+rewrite H0 in H.
+apply H.
+Qed.
+
+Lemma cPairInj2 : forall a b c d : nat, cPair a b = cPair c d -> b = d.
+Proof.
+intros.
+assert (a + b = c + d).
+apply cPairInjHelp.
+auto.
+assert (a = c).
+eapply cPairInj1.
+apply H.
+eapply plus_reg_l.
+rewrite H1 in H0.
+apply H0.
+Qed.
+
+End CPair_Injectivity.
+
+Section CPair_projections.
+
+Let searchXY (a : nat) :=
+ boundedSearch (fun a y : nat => ltBool a (sumToN (S y))) a.
+
+Definition cPairPi1 (a : nat) := a - sumToN (searchXY a).
+Definition cPairPi2 (a : nat) := searchXY a - cPairPi1 a.
+
+Lemma cPairProjectionsHelp :
+ forall a b : nat, b < sumToN (S a) -> sumToN a <= b -> searchXY b = a.
+Proof.
+intros.
+unfold searchXY in |- *.
+induction (boundedSearch2 (fun b y : nat => ltBool b (sumToN (S y))) b).
+rewrite H1.
+induction (eq_nat_dec b a).
+auto.
+elim (ltBoolFalse b (sumToN (S a))).
+apply (boundedSearch1 (fun b y : nat => ltBool b (sumToN (S y))) b).
+rewrite H1.
+induction (nat_total_order _ _ b0).
+elim (lt_not_le _ _ H2).
+apply le_trans with (sumToN a).
+apply sumToN1.
+auto.
+auto.
+auto.
+set (c := boundedSearch (fun b y : nat => ltBool b (sumToN (S y))) b) in *.
+induction (eq_nat_dec c a).
+auto.
+elim (ltBoolFalse b (sumToN (S a))).
+apply (boundedSearch1 (fun b y : nat => ltBool b (sumToN (S y))) b).
+fold c in |- *.
+induction (nat_total_order _ _ b0).
+elim (le_not_lt _ _ H0).
+apply lt_le_trans with (sumToN (S c)).
+apply ltBoolTrue.
+auto.
+assert (S c <= a).
+apply lt_n_Sm_le.
+apply lt_n_S.
+auto.
+apply sumToN2.
+auto.
+auto.
+auto.
+Qed.
+
+Lemma cPairProjections : forall a : nat, cPair (cPairPi1 a) (cPairPi2 a) = a.
+Proof.
+assert
+ (forall a b : nat, b < sumToN a -> cPair (cPairPi1 b) (cPairPi2 b) = b).
+intros.
+induction a as [| a Hreca].
+simpl in H.
+elim (lt_n_O _ H).
+induction (le_or_lt (sumToN a) b).
+assert (searchXY b = a).
+apply cPairProjectionsHelp; auto.
+unfold cPair in |- *.
+replace (cPairPi1 b + cPairPi2 b) with a.
+unfold cPairPi1 in |- *.
+rewrite H1.
+rewrite plus_comm.
+rewrite <- le_plus_minus.
+reflexivity.
+auto.
+unfold cPairPi2 in |- *.
+rewrite <- le_plus_minus.
+auto.
+unfold cPairPi1 in |- *.
+rewrite H1.
+simpl in H.
+apply (fun p n m : nat => plus_le_reg_l n m p) with (sumToN a).
+rewrite <- le_plus_minus.
+rewrite plus_comm.
+apply lt_n_Sm_le.
+auto.
+auto.
+apply Hreca.
+auto.
+intros.
+apply H with (S a).
+apply lt_le_trans with (S a).
+apply lt_n_Sn.
+apply sumToN1.
+Qed.
+
+Lemma cPairProjections1 : forall a b : nat, cPairPi1 (cPair a b) = a.
+Proof.
+intros.
+unfold cPair in |- *.
+unfold cPairPi1 in |- *.
+replace (searchXY (a + sumToN (a + b))) with (a + b).
+rewrite plus_comm.
+apply minus_plus.
+symmetry in |- *.
+apply cPairProjectionsHelp.
+simpl in |- *.
+apply le_lt_n_Sm.
+apply plus_le_compat_r.
+apply le_plus_l.
+apply le_plus_r.
+Qed.
+
+Lemma cPairProjections2 : forall a b : nat, cPairPi2 (cPair a b) = b.
+Proof.
+intros.
+unfold cPairPi2 in |- *.
+rewrite cPairProjections1.
+unfold cPair in |- *.
+replace (searchXY (a + sumToN (a + b))) with (a + b).
+apply minus_plus.
+symmetry in |- *.
+apply cPairProjectionsHelp.
+simpl in |- *.
+apply le_lt_n_Sm.
+apply plus_le_compat_r.
+apply le_plus_l.
+apply le_plus_r.
+Qed.
+
+End CPair_projections.
+
+Section CPair_Order.
+
+Lemma cPairLe1 : forall a b : nat, a <= cPair a b.
+Proof.
+intros.
+unfold cPair in |- *.
+apply le_plus_l.
+Qed.
+
+Lemma cPairLe1A : forall a : nat, cPairPi1 a <= a.
+intros.
+apply le_trans with (cPair (cPairPi1 a) (cPairPi2 a)).
+apply cPairLe1.
+rewrite cPairProjections.
+apply le_n.
+Qed.
+
+Lemma cPairLe2 : forall a b : nat, b <= cPair a b.
+Proof.
+intros.
+unfold cPair in |- *.
+eapply le_trans.
+apply le_plus_r.
+apply plus_le_compat_l.
+apply le_trans with (a + b).
+apply le_plus_r.
+apply sumToN1.
+Qed.
+
+Lemma cPairLe2A : forall a : nat, cPairPi2 a <= a.
+intros.
+apply le_trans with (cPair (cPairPi1 a) (cPairPi2 a)).
+apply cPairLe2.
+rewrite cPairProjections.
+apply le_n.
+Qed.
+
+Lemma cPairLe3 :
+ forall a b c d : nat, a <= b -> c <= d -> cPair a c <= cPair b d.
+Proof.
+intros.
+unfold cPair in |- *.
+apply le_trans with (a + sumToN (b + d)).
+apply plus_le_compat_l.
+apply sumToN2.
+apply le_trans with (a + d).
+apply plus_le_compat_l.
+auto.
+apply plus_le_compat_r.
+auto.
+apply plus_le_compat_r.
+auto.
+Qed.
+
+Lemma cPairLt1 : forall a b : nat, a < cPair a (S b).
+Proof.
+intros.
+unfold cPair in |- *.
+rewrite (plus_comm a (S b)).
+simpl in |- *.
+rewrite plus_comm.
+simpl in |- *.
+rewrite plus_comm.
+unfold lt in |- *.
+apply le_n_S.
+apply le_plus_l.
+Qed.
+
+Lemma cPairLt2 : forall a b : nat, b < cPair (S a) b.
+Proof.
+intros.
+unfold cPair in |- *.
+simpl in |- *.
+unfold lt in |- *.
+apply le_n_S.
+eapply le_trans.
+apply le_plus_r.
+apply plus_le_compat_l.
+apply le_S.
+eapply le_trans.
+apply le_plus_l.
+rewrite plus_comm.
+apply le_plus_l.
+Qed.
+
+End CPair_Order.
diff --git a/boolean_completeness/pairing/extEqualNat.v b/boolean_completeness/pairing/extEqualNat.v
new file mode 100644
index 0000000..e7562cb
--- /dev/null
+++ b/boolean_completeness/pairing/extEqualNat.v
@@ -0,0 +1,76 @@
+(** This code is part of Russell O'Connor's formalisation of Goedel's
+ Incompleteness Theorem in Coq, released into the Public Domain.
+
+ http://r6.ca/Goedel/goedel1.html
+*)
+Require Import Arith.
+
+Fixpoint naryFunc (n : nat) : Set :=
+ match n with
+ | O => nat
+ | S n => nat -> naryFunc n
+ end.
+
+Fixpoint naryRel (n : nat) : Set :=
+ match n with
+ | O => bool
+ | S n => nat -> naryRel n
+ end.
+
+Definition extEqual (n : nat) (a b : naryFunc n) : Prop.
+intros.
+induction n as [| n Hrecn].
+exact (a = b).
+exact (forall c : nat, Hrecn (a c) (b c)).
+Defined.
+
+Lemma extEqualRefl : forall (n : nat) (a : naryFunc n), extEqual n a a.
+Proof.
+intros.
+induction n as [| n Hrecn].
+simpl in |- *.
+reflexivity.
+simpl in |- *.
+intro.
+apply Hrecn.
+Qed.
+
+Lemma extEqualSym :
+ forall (n : nat) (a b : naryFunc n), extEqual n a b -> extEqual n b a.
+Proof.
+intros.
+induction n as [| n Hrecn].
+simpl in |- *.
+symmetry in |- *.
+apply H.
+simpl in |- *.
+intros.
+apply Hrecn.
+simpl in H.
+apply H.
+Qed.
+
+Lemma extEqualTrans :
+ forall (n : nat) (a b c : naryFunc n),
+ extEqual n a b -> extEqual n b c -> extEqual n a c.
+Proof.
+intros.
+induction n as [| n Hrecn].
+simpl in |- *.
+transitivity b; auto.
+simpl in |- *.
+intros.
+eapply Hrecn.
+simpl in H.
+apply (H c0).
+apply (H0 c0).
+Qed.
+
+Fixpoint charFunction (n : nat) : naryRel n -> naryFunc n :=
+ match n return (naryRel n -> naryFunc n) with
+ | O => fun R : bool => match R with
+ | true => 1
+ | false => 0
+ end
+ | S m => fun (R : naryRel (S m)) (a : nat) => charFunction m (R a)
+ end.
\ No newline at end of file
diff --git a/boolean_completeness/pairing/misc.v b/boolean_completeness/pairing/misc.v
new file mode 100644
index 0000000..5601fba
--- /dev/null
+++ b/boolean_completeness/pairing/misc.v
@@ -0,0 +1,42 @@
+(** This code is part of Russell O'Connor's formalisation of Goedel's
+ Incompleteness Theorem in Coq, released into the Public Domain.
+
+ http://r6.ca/Goedel/goedel1.html
+*)
+Require Import Eqdep_dec.
+
+Lemma inj_right_pair2 :
+ forall A : Set,
+ (forall x y : A, {x = y} + {x <> y}) ->
+ forall (x : A) (P : A -> Set) (y y' : P x),
+ existS P x y = existS P x y' -> y = y'.
+Proof.
+intros.
+set
+ (Proj :=
+ fun (e : sigS P) (def : P x) =>
+ match e with
+ | existS x' dep =>
+ match H x' x with
+ | left eqdep => eq_rec x' P dep x eqdep
+ | _ => def
+ end
+ end) in *.
+cut (Proj (existS P x y) y = Proj (existS P x y') y).
+simpl in |- *.
+induction (H x x).
+intro e.
+set
+ (B :=
+ K_dec_set H
+ (fun z : x = x =>
+ eq_rec x P y x z = eq_rec x P y' x z ->
+ eq_rec x P y x (refl_equal x) = eq_rec x P y' x (refl_equal x))
+ (fun z : eq_rec x P y x (refl_equal x) = eq_rec x P y' x (refl_equal x) =>
+ z) a e) in *.
+apply B.
+elim b.
+auto.
+rewrite H0.
+reflexivity.
+Qed.
\ No newline at end of file
diff --git a/boolean_completeness/pairing/primRec.v b/boolean_completeness/pairing/primRec.v
new file mode 100644
index 0000000..24801bc
--- /dev/null
+++ b/boolean_completeness/pairing/primRec.v
@@ -0,0 +1,2350 @@
+(** This code is part of Russell O'Connor's formalisation of Goedel's
+ Incompleteness Theorem in Coq, released into the Public Domain.
+
+ http://r6.ca/Goedel/goedel1.html
+*)
+Require Import Arith.
+Require Import Peano_dec.
+Require Import Compare_dec.
+Require Import Coq.Lists.List.
+Require Import Eqdep_dec.
+Require Import extEqualNat.
+Require Import Bvector.
+Require Import misc.
+Require Export Bool.
+Require Export EqNat.
+Require Import Even.
+Require Import Max.
+
+Inductive PrimRec : nat -> Set :=
+ | succFunc : PrimRec 1
+ | zeroFunc : PrimRec 0
+ | projFunc : forall n m : nat, m < n -> PrimRec n
+ | composeFunc :
+ forall (n m : nat) (g : PrimRecs n m) (h : PrimRec m), PrimRec n
+ | primRecFunc :
+ forall (n : nat) (g : PrimRec n) (h : PrimRec (S (S n))), PrimRec (S n)
+with PrimRecs : nat -> nat -> Set :=
+ | PRnil : forall n : nat, PrimRecs n 0
+ | PRcons : forall n m : nat, PrimRec n -> PrimRecs n m -> PrimRecs n (S m).
+
+Scheme PrimRec_PrimRecs_rec := Induction for PrimRec
+ Sort Set
+ with PrimRecs_PrimRec_rec := Induction for PrimRecs
+ Sort Set.
+
+Scheme PrimRec_PrimRecs_ind := Induction for PrimRec
+ Sort Prop
+ with PrimRecs_PrimRec_ind := Induction for PrimRecs
+ Sort Prop.
+
+Fixpoint evalConstFunc (n m : nat) {struct n} : naryFunc n :=
+ match n return (naryFunc n) with
+ | O => m
+ | S n' => fun _ => evalConstFunc n' m
+ end.
+
+(* The parameters are number in opposite order.
+ So proj(2,0)(a,b) = b. *)
+Fixpoint evalProjFunc (n : nat) : forall m : nat, m < n -> naryFunc n :=
+ match n return (forall m : nat, m < n -> naryFunc n) with
+ | O => fun (m : nat) (l : m < 0) => False_rec _ (lt_n_O _ l)
+ | S n' =>
+ fun (m : nat) (l : m < S n') =>
+ match eq_nat_dec m n' with
+ | left _ => fun a : nat => evalConstFunc _ a
+ | right l1 =>
+ fun _ =>
+ evalProjFunc n' m
+ match le_lt_or_eq _ _ (lt_n_Sm_le _ _ l) with
+ | or_introl l2 => l2
+ | or_intror l2 => False_ind _ (l1 l2)
+ end
+ end
+ end.
+
+Lemma evalProjFuncInd :
+ forall (n m : nat) (p1 p2 : m < n),
+ evalProjFunc n m p1 = evalProjFunc n m p2.
+Proof.
+intro.
+induction n as [| n Hrecn].
+intros.
+elim (lt_n_O _ p1).
+intros.
+simpl in |- *.
+induction (eq_nat_dec m n).
+reflexivity.
+rewrite
+ (Hrecn _
+ match le_lt_or_eq m n (lt_n_Sm_le m n p1) with
+ | or_introl l2 => l2
+ | or_intror l2 => False_ind (m < n) (b l2)
+ end
+ match le_lt_or_eq m n (lt_n_Sm_le m n p2) with
+ | or_introl l2 => l2
+ | or_intror l2 => False_ind (m < n) (b l2)
+ end).
+reflexivity.
+Qed.
+
+Fixpoint evalList (m : nat) (l : Vector.t nat m) {struct l} :
+ naryFunc m -> nat :=
+ match l with
+ | Vector.nil => fun x : naryFunc 0 => x
+ | Vector.cons a n l' => fun x : naryFunc (S n) => evalList n l' (x a)
+ end.
+
+Fixpoint evalOneParamList (n m a : nat) (l : Vector.t (naryFunc (S n)) m)
+ {struct l} : Vector.t (naryFunc n) m :=
+ match l in (Vector.t _ m) return (Vector.t (naryFunc n) m) with
+ | Vector.nil => Vector.nil (naryFunc n)
+ | Vector.cons f m' l' => Vector.cons _ (f a) m' (evalOneParamList n m' a l')
+ end.
+
+Fixpoint evalComposeFunc (n : nat) :
+ forall m : nat, Vector.t (naryFunc n) m -> naryFunc m -> naryFunc n :=
+ match
+ n
+ return
+ (forall m : nat, Vector.t (naryFunc n) m -> naryFunc m -> naryFunc n)
+ with
+ | O => evalList
+ | S n' =>
+ fun (m : nat) (l : Vector.t (naryFunc (S n')) m)
+ (f : naryFunc m) (a : nat) =>
+ evalComposeFunc n' m (evalOneParamList _ _ a l) f
+ end.
+
+Fixpoint compose2 (n : nat) : naryFunc n -> naryFunc (S n) -> naryFunc n :=
+ match n return (naryFunc n -> naryFunc (S n) -> naryFunc n) with
+ | O => fun (a : nat) (g : nat -> nat) => g a
+ | S n' =>
+ fun (f : naryFunc (S n')) (g : naryFunc (S (S n'))) (a : nat) =>
+ compose2 n' (f a) (fun x : nat => g x a)
+ end.
+
+Fixpoint evalPrimRecFunc (n : nat) (g : naryFunc n)
+ (h : naryFunc (S (S n))) (a : nat) {struct a} : naryFunc n :=
+ match a with
+ | O => g
+ | S a' => compose2 _ (evalPrimRecFunc n g h a') (h a')
+ end.
+
+Fixpoint evalPrimRec (n : nat) (f : PrimRec n) {struct f} :
+ naryFunc n :=
+ match f in (PrimRec n) return (naryFunc n) with
+ | succFunc => S
+ | zeroFunc => 0
+ | projFunc n m pf => evalProjFunc n m pf
+ | composeFunc n m l f =>
+ evalComposeFunc n m (evalPrimRecs _ _ l) (evalPrimRec _ f)
+ | primRecFunc n g h =>
+ evalPrimRecFunc n (evalPrimRec _ g) (evalPrimRec _ h)
+ end
+
+ with evalPrimRecs (n m : nat) (fs : PrimRecs n m) {struct fs} :
+ Vector.t (naryFunc n) m :=
+ match fs in (PrimRecs n m) return (Vector.t (naryFunc n) m) with
+ | PRnil a => Vector.nil (naryFunc a)
+ | PRcons a b g gs => Vector.cons _ (evalPrimRec _ g) _ (evalPrimRecs _ _ gs)
+ end.
+
+Definition extEqualVectorGeneral : forall (n m : nat) (l : Vector.t (naryFunc n) m)
+ (m' : nat) (l' : Vector.t (naryFunc n) m') , Prop.
+intros n m l.
+induction l as [| a n0 l Hrecl].
+intros.
+destruct l' as [| a n0 v].
+exact True.
+exact False.
+intros.
+destruct l' as [| a0 n1 v].
+exact False.
+exact (extEqual n a a0 /\ Hrecl _ v).
+Defined.
+
+Definition extEqualVector : forall (n m : nat) (l l' : Vector.t (naryFunc n) m) , Prop.
+intros.
+eapply extEqualVectorGeneral.
+apply l.
+apply l'.
+Defined.
+
+Lemma extEqualVectorRefl :
+ forall (n m : nat) (l : Vector.t (naryFunc n) m), extEqualVector n m l l.
+Proof.
+intros.
+unfold extEqualVector in |- *.
+induction l as [| a n0 l Hrecl].
+simpl in |- *.
+auto.
+simpl in |- *.
+split.
+apply extEqualRefl.
+auto.
+Qed.
+
+Lemma extEqualOneParamList :
+ forall (n m : nat) (l1 l2 : Vector.t (naryFunc (S n)) m) (c : nat),
+ extEqualVector (S n) m l1 l2 ->
+ extEqualVector n m (evalOneParamList n m c l1) (evalOneParamList n m c l2).
+Proof.
+unfold extEqualVector in |- *.
+assert
+ (forall (n m : nat) (l1 : Vector.t (naryFunc (S n)) m)
+ (m' : nat) (l2 : Vector.t (naryFunc (S n)) m') (c : nat),
+ extEqualVectorGeneral (S n) m l1 m' l2 ->
+ extEqualVectorGeneral n m (evalOneParamList n m c l1) m'
+ (evalOneParamList n m' c l2)).
+intros n m l1.
+induction l1 as [| a n0 l1 Hrecl1].
+intros.
+simpl in |- *.
+induction l2 as [| a n0 l2 Hrecl2].
+simpl in |- *.
+auto.
+simpl in H.
+elim H.
+intros.
+induction l2 as [| a0 n1 l2 Hrecl2].
+simpl in H.
+elim H.
+simpl in |- *.
+simpl in H.
+induction H as (H, H0).
+auto.
+intros.
+apply H.
+auto.
+Qed.
+
+Lemma extEqualCompose :
+ forall (n m : nat) (l1 l2 : Vector.t (naryFunc n) m) (f1 f2 : naryFunc m),
+ extEqualVector n m l1 l2 ->
+ extEqual m f1 f2 ->
+ extEqual n (evalComposeFunc n m l1 f1) (evalComposeFunc n m l2 f2).
+Proof.
+intro.
+induction n as [| n Hrecn]; simpl in |- *.
+intros m l1.
+induction l1 as [| a n l1 Hrecl1]; simpl in |- *.
+intros.
+rewrite H0.
+clear H0.
+(*
+Print evalList.
+Fixpoint evalList (m : nat) (l : Vector.t nat m) {struct l} :
+ naryFunc m -> nat :=
+ match l with
+ | Vector.nil => fun x : naryFunc 0 => x
+ | Vector.cons a n l' => fun x : naryFunc (S n) => evalList n l' (x a)
+ end.
+*)
+
+Definition Vhead {T:Type} (n:nat) (v:Vector.t T (S n)) :=
+ match v with
+ | Vector.cons a _ _ => a
+ end.
+
+Definition Vtail {T:Type} (n:nat) (v:Vector.t T (S n)) :=
+ match v with
+ | Vector.cons _ _ v => v
+ end.
+
+Lemma V0_eq : forall (T:Type) (v : Vector.t T 0), v = Vector.nil T.
+Proof.
+ intros T v.
+ Print Vector.t.
+ refine (match v with
+ | Vector.nil => _
+ | Vector.cons a n v' => _
+ end).
+ reflexivity.
+ exact (fun A x => x).
+Qed.
+
+Lemma VSn_eq :
+ forall (T:Type) (n : nat) (v : Vector.t T (S n)), v = Vector.cons T (Vhead _ v) _ (Vtail _ v).
+Proof.
+ intros T n v.
+ refine (match v with
+ | Vector.nil => _
+ | Vector.cons a n v' => _
+ end).
+ exact (fun A x => x).
+ reflexivity.
+Qed.
+
+rewrite (V0_eq _ l2).
+simpl in |- *.
+reflexivity.
+intros.
+rewrite VSn_eq in H.
+unfold extEqualVector in H.
+simpl in H.
+induction H as (H, H1).
+rewrite (VSn_eq _ _ l2).
+simpl in |- *.
+apply Hrecl1.
+apply H1.
+rewrite H.
+apply H0.
+intros.
+apply Hrecn.
+apply extEqualOneParamList.
+auto.
+auto.
+Qed.
+
+Lemma extEqualCompose2 :
+ forall (n : nat) (f1 f2 : naryFunc n),
+ extEqual n f1 f2 ->
+ forall g1 g2 : naryFunc (S n),
+ extEqual (S n) g1 g2 -> extEqual n (compose2 n f1 g1) (compose2 n f2 g2).
+Proof.
+intro.
+induction n as [| n Hrecn]; simpl in |- *; intros.
+rewrite H.
+apply H0.
+apply Hrecn; simpl in |- *; intros; auto.
+Qed.
+
+Lemma extEqualPrimRec :
+ forall (n : nat) (g1 g2 : naryFunc n) (h1 h2 : naryFunc (S (S n))),
+ extEqual n g1 g2 ->
+ extEqual (S (S n)) h1 h2 ->
+ extEqual (S n) (evalPrimRecFunc n g1 h1) (evalPrimRecFunc n g2 h2).
+Proof.
+intros.
+simpl in |- *.
+intros.
+induction c as [| c Hrecc].
+simpl in |- *.
+auto.
+simpl in |- *.
+cut (extEqual n (evalPrimRecFunc n g1 h1 c) (evalPrimRecFunc n g2 h2 c)).
+cut (extEqual (S n) (h1 c) (h2 c)).
+generalize (h1 c) (h2 c) (evalPrimRecFunc n g1 h1 c)
+ (evalPrimRecFunc n g2 h2 c).
+fold (naryFunc (S n)) in |- *.
+clear Hrecc H0 h1 h2 g1 g2 H.
+induction n as [| n Hrecn].
+simpl in |- *.
+intros.
+rewrite H0.
+apply H.
+simpl in |- *.
+intros.
+apply Hrecn.
+simpl in |- *.
+intros.
+apply H.
+apply H0.
+simpl in |- *.
+intros.
+simpl in H0.
+apply H0.
+auto.
+Qed.
+
+Definition isPR (n : nat) (f : naryFunc n) : Set :=
+ {p : PrimRec n | extEqual n (evalPrimRec _ p) f}.
+
+Definition isPRrel (n : nat) (R : naryRel n) : Set :=
+ isPR n (charFunction n R).
+
+Lemma succIsPR : isPR 1 S.
+Proof.
+exists succFunc.
+simpl in |- *.
+auto.
+Qed.
+
+Lemma const0_NIsPR : forall n : nat, isPR 0 n.
+Proof.
+simple induction n.
+exists zeroFunc.
+simpl in |- *.
+reflexivity.
+intros.
+induction H as (x, p).
+exists (composeFunc _ _ (PRcons _ _ x (PRnil _)) succFunc).
+simpl in |- *.
+simpl in p.
+rewrite p.
+reflexivity.
+Qed.
+
+Lemma const1_NIsPR : forall n : nat, isPR 1 (fun _ => n).
+Proof.
+intros.
+assert (isPR 0 n).
+apply const0_NIsPR.
+induction H as (x, p).
+exists (composeFunc 1 _ (PRnil _) x).
+simpl in |- *.
+simpl in p.
+auto.
+Qed.
+
+Lemma idIsPR : isPR 1 (fun x : nat => x).
+Proof.
+assert (0 < 1).
+auto.
+exists (projFunc _ _ H).
+simpl in |- *.
+auto.
+Qed.
+
+Lemma pi1_2IsPR : isPR 2 (fun a b : nat => a).
+Proof.
+assert (1 < 2).
+auto.
+exists (projFunc _ _ H).
+simpl in |- *.
+auto.
+Qed.
+
+Lemma pi2_2IsPR : isPR 2 (fun a b : nat => b).
+Proof.
+assert (0 < 2).
+auto.
+exists (projFunc _ _ H).
+simpl in |- *.
+auto.
+Qed.
+
+Lemma pi1_3IsPR : isPR 3 (fun a b c : nat => a).
+Proof.
+assert (2 < 3).
+auto.
+exists (projFunc _ _ H).
+simpl in |- *.
+auto.
+Qed.
+
+Lemma pi2_3IsPR : isPR 3 (fun a b c : nat => b).
+Proof.
+assert (1 < 3).
+auto.
+exists (projFunc _ _ H).
+simpl in |- *.
+auto.
+Qed.
+
+Lemma pi3_3IsPR : isPR 3 (fun a b c : nat => c).
+Proof.
+assert (0 < 3).
+auto.
+exists (projFunc _ _ H).
+simpl in |- *.
+auto.
+Qed.
+
+Lemma pi1_4IsPR : isPR 4 (fun a b c d : nat => a).
+Proof.
+assert (3 < 4).
+auto.
+exists (projFunc _ _ H).
+simpl in |- *.
+auto.
+Qed.
+
+Lemma pi2_4IsPR : isPR 4 (fun a b c d : nat => b).
+Proof.
+assert (2 < 4).
+auto.
+exists (projFunc _ _ H).
+simpl in |- *.
+auto.
+Qed.
+
+Lemma pi3_4IsPR : isPR 4 (fun a b c d : nat => c).
+Proof.
+assert (1 < 4).
+auto.
+exists (projFunc _ _ H).
+simpl in |- *.
+auto.
+Qed.
+
+Lemma pi4_4IsPR : isPR 4 (fun a b c d : nat => d).
+Proof.
+assert (0 < 4).
+auto.
+exists (projFunc _ _ H).
+simpl in |- *.
+auto.
+Qed.
+
+Lemma filter01IsPR :
+ forall g : nat -> nat, isPR 1 g -> isPR 2 (fun a b : nat => g b).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 2 (fun a b : nat => b)).
+apply pi2_2IsPR.
+induction H as (x0, p0).
+simpl in p0.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRnil _)) x).
+simpl in |- *.
+intros.
+replace (g c0) with (g (evalPrimRec 2 x0 c c0)).
+rewrite <- p.
+auto.
+rewrite p0.
+auto.
+Qed.
+
+Lemma filter10IsPR :
+ forall g : nat -> nat, isPR 1 g -> isPR 2 (fun a b : nat => g a).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 2 (fun a b : nat => a)).
+apply pi1_2IsPR.
+induction H as (x0, p0).
+simpl in p0.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRnil _)) x).
+simpl in |- *.
+intros.
+replace (g c) with (g (evalPrimRec 2 x0 c c0)).
+rewrite <- p.
+auto.
+rewrite p0.
+auto.
+Qed.
+
+Lemma filter100IsPR :
+ forall g : nat -> nat, isPR 1 g -> isPR 3 (fun a b c : nat => g a).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 3 (fun a b c : nat => a)).
+apply pi1_3IsPR.
+induction H as (x0, p0).
+simpl in p0.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRnil _)) x).
+simpl in |- *.
+intros.
+replace (g c) with (g (evalPrimRec 3 x0 c c0 c1)).
+rewrite <- p.
+auto.
+rewrite p0.
+auto.
+Qed.
+
+Lemma filter010IsPR :
+ forall g : nat -> nat, isPR 1 g -> isPR 3 (fun a b c : nat => g b).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 3 (fun a b c : nat => b)).
+apply pi2_3IsPR.
+induction H as (x0, p0).
+simpl in p0.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRnil _)) x).
+simpl in |- *.
+intros.
+replace (g c0) with (g (evalPrimRec 3 x0 c c0 c1)).
+rewrite <- p.
+auto.
+rewrite p0.
+auto.
+Qed.
+
+Lemma filter001IsPR :
+ forall g : nat -> nat, isPR 1 g -> isPR 3 (fun a b c : nat => g c).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 3 (fun a b c : nat => c)).
+apply pi3_3IsPR.
+induction H as (x0, p0).
+simpl in p0.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRnil _)) x).
+simpl in |- *.
+intros.
+replace (g c1) with (g (evalPrimRec 3 x0 c c0 c1)).
+rewrite <- p.
+auto.
+rewrite p0.
+auto.
+Qed.
+
+Lemma filter011IsPR :
+ forall g : nat -> nat -> nat, isPR 2 g -> isPR 3 (fun a b c : nat => g b c).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 3 (fun a b c : nat => b)).
+apply pi2_3IsPR.
+induction H as (x0, p0).
+simpl in p0.
+assert (isPR 3 (fun a b c : nat => c)).
+apply pi3_3IsPR.
+induction H as (x1, p1).
+simpl in p1.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRcons _ _ x1 (PRnil _))) x).
+simpl in |- *.
+intros.
+replace (g c0 c1) with
+ (g (evalPrimRec 3 x0 c c0 c1) (evalPrimRec 3 x1 c c0 c1)).
+rewrite <- p.
+auto.
+rewrite p0.
+rewrite p1.
+auto.
+Qed.
+
+Lemma filter110IsPR :
+ forall g : nat -> nat -> nat, isPR 2 g -> isPR 3 (fun a b c : nat => g a b).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 3 (fun a b c : nat => a)).
+apply pi1_3IsPR.
+induction H as (x0, p0).
+simpl in p0.
+assert (isPR 3 (fun a b c : nat => b)).
+apply pi2_3IsPR.
+induction H as (x1, p1).
+simpl in p1.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRcons _ _ x1 (PRnil _))) x).
+simpl in |- *.
+intros.
+replace (g c c0) with
+ (g (evalPrimRec 3 x0 c c0 c1) (evalPrimRec 3 x1 c c0 c1)).
+rewrite <- p.
+auto.
+rewrite p0.
+rewrite p1.
+auto.
+Qed.
+
+Lemma filter101IsPR :
+ forall g : nat -> nat -> nat, isPR 2 g -> isPR 3 (fun a b c : nat => g a c).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 3 (fun a b c : nat => a)).
+apply pi1_3IsPR.
+induction H as (x0, p0).
+simpl in p0.
+assert (isPR 3 (fun a b c : nat => c)).
+apply pi3_3IsPR.
+induction H as (x1, p1).
+simpl in p1.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRcons _ _ x1 (PRnil _))) x).
+simpl in |- *.
+intros.
+replace (g c c1) with
+ (g (evalPrimRec 3 x0 c c0 c1) (evalPrimRec 3 x1 c c0 c1)).
+rewrite <- p.
+auto.
+rewrite p0.
+rewrite p1.
+auto.
+Qed.
+
+Lemma filter0011IsPR :
+ forall g : nat -> nat -> nat,
+ isPR 2 g -> isPR 4 (fun a b c d : nat => g c d).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 4 (fun a b c d : nat => c)).
+apply pi3_4IsPR.
+induction H as (x0, p0).
+simpl in p0.
+assert (isPR 4 (fun a b c d : nat => d)).
+apply pi4_4IsPR.
+induction H as (x1, p1).
+simpl in p1.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRcons _ _ x1 (PRnil _))) x).
+simpl in |- *.
+intros.
+replace (g c1 c2) with
+ (g (evalPrimRec 4 x0 c c0 c1 c2) (evalPrimRec 4 x1 c c0 c1 c2)).
+rewrite <- p.
+auto.
+rewrite p0.
+rewrite p1.
+auto.
+Qed.
+
+Lemma filter1000IsPR :
+ forall g : nat -> nat, isPR 1 g -> isPR 4 (fun a b c d : nat => g a).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 4 (fun a b c d : nat => a)).
+apply pi1_4IsPR.
+induction H as (x0, p0).
+simpl in p0.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRnil _)) x).
+simpl in |- *.
+intros.
+replace (g c) with (g (evalPrimRec 4 x0 c c0 c1 c2)).
+rewrite <- p.
+auto.
+rewrite p0.
+auto.
+Qed.
+
+Lemma filter1011IsPR :
+ forall g : nat -> nat -> nat -> nat,
+ isPR 3 g -> isPR 4 (fun a b c d : nat => g a c d).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 4 (fun a b c d : nat => a)).
+apply pi1_4IsPR.
+assert (isPR 4 (fun a b c d : nat => c)).
+apply pi3_4IsPR.
+assert (isPR 4 (fun a b c d : nat => d)).
+apply pi4_4IsPR.
+induction H as (x0, p0).
+simpl in p0.
+induction H0 as (x1, p1).
+simpl in p1.
+induction H1 as (x2, p2).
+simpl in p2.
+exists
+ (composeFunc _ _ (PRcons _ _ x0 (PRcons _ _ x1 (PRcons _ _ x2 (PRnil _)))) x).
+simpl in |- *.
+intros.
+replace (g c c1 c2) with
+ (g (evalPrimRec 4 x0 c c0 c1 c2) (evalPrimRec 4 x1 c c0 c1 c2)
+ (evalPrimRec 4 x2 c c0 c1 c2)).
+rewrite <- p.
+auto.
+rewrite p0.
+auto.
+Qed.
+
+Lemma filter1100IsPR :
+ forall g : nat -> nat -> nat,
+ isPR 2 g -> isPR 4 (fun a b c d : nat => g a b).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+assert (isPR 4 (fun a b c d : nat => a)).
+apply pi1_4IsPR.
+assert (isPR 4 (fun a b c d : nat => b)).
+apply pi2_4IsPR.
+induction H as (x0, p0).
+simpl in p0.
+induction H0 as (x1, p1).
+simpl in p1.
+exists (composeFunc _ _ (PRcons _ _ x0 (PRcons _ _ x1 (PRnil _))) x).
+simpl in |- *.
+intros.
+replace (g c c0) with
+ (g (evalPrimRec 4 x0 c c0 c1 c2) (evalPrimRec 4 x1 c c0 c1 c2)).
+rewrite <- p.
+auto.
+auto.
+Qed.
+
+Lemma compose1_1IsPR :
+ forall f : nat -> nat,
+ isPR 1 f ->
+ forall g : nat -> nat, isPR 1 g -> isPR 1 (fun x : nat => g (f x)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+exists (composeFunc _ _ (PRcons _ _ x (PRnil _)) x0).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+auto.
+Qed.
+
+Lemma compose1_2IsPR :
+ forall f : nat -> nat,
+ isPR 1 f ->
+ forall f' : nat -> nat,
+ isPR 1 f' ->
+ forall g : nat -> nat -> nat,
+ isPR 2 g -> isPR 1 (fun x : nat => g (f x) (f' x)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+induction H1 as (x1, p1); simpl in p1.
+exists (composeFunc _ _ (PRcons _ _ x (PRcons _ _ x0 (PRnil _))) x1).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+rewrite <- p1.
+auto.
+Qed.
+
+Lemma compose1_3IsPR :
+ forall f1 : nat -> nat,
+ isPR 1 f1 ->
+ forall f2 : nat -> nat,
+ isPR 1 f2 ->
+ forall f3 : nat -> nat,
+ isPR 1 f3 ->
+ forall g : nat -> nat -> nat -> nat,
+ isPR 3 g -> isPR 1 (fun x : nat => g (f1 x) (f2 x) (f3 x)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+induction H1 as (x1, p1); simpl in p1.
+induction H2 as (x2, p2); simpl in p2.
+exists
+ (composeFunc _ _ (PRcons _ _ x (PRcons _ _ x0 (PRcons _ _ x1 (PRnil _)))) x2).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+rewrite <- p1.
+rewrite <- p2.
+auto.
+Qed.
+
+Lemma compose2_1IsPR :
+ forall f : nat -> nat -> nat,
+ isPR 2 f ->
+ forall g : nat -> nat, isPR 1 g -> isPR 2 (fun x y : nat => g (f x y)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+exists (composeFunc _ _ (PRcons _ _ x (PRnil _)) x0).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+auto.
+Qed.
+
+Lemma compose2_2IsPR :
+ forall f : nat -> nat -> nat,
+ isPR 2 f ->
+ forall g : nat -> nat -> nat,
+ isPR 2 g ->
+ forall h : nat -> nat -> nat,
+ isPR 2 h -> isPR 2 (fun x y : nat => h (f x y) (g x y)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+induction H1 as (x1, p1); simpl in p1.
+exists (composeFunc _ _ (PRcons _ _ x (PRcons _ _ x0 (PRnil _))) x1).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+rewrite <- p1.
+auto.
+Qed.
+
+Lemma compose2_3IsPR :
+ forall f1 : nat -> nat -> nat,
+ isPR 2 f1 ->
+ forall f2 : nat -> nat -> nat,
+ isPR 2 f2 ->
+ forall f3 : nat -> nat -> nat,
+ isPR 2 f3 ->
+ forall g : nat -> nat -> nat -> nat,
+ isPR 3 g -> isPR 2 (fun x y : nat => g (f1 x y) (f2 x y) (f3 x y)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+induction H1 as (x1, p1); simpl in p1.
+induction H2 as (x2, p2); simpl in p2.
+exists
+ (composeFunc _ _ (PRcons _ _ x (PRcons _ _ x0 (PRcons _ _ x1 (PRnil _)))) x2).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+rewrite <- p1.
+rewrite <- p2.
+auto.
+Qed.
+
+Lemma compose2_4IsPR :
+ forall f1 : nat -> nat -> nat,
+ isPR 2 f1 ->
+ forall f2 : nat -> nat -> nat,
+ isPR 2 f2 ->
+ forall f3 : nat -> nat -> nat,
+ isPR 2 f3 ->
+ forall f4 : nat -> nat -> nat,
+ isPR 2 f4 ->
+ forall g : nat -> nat -> nat -> nat -> nat,
+ isPR 4 g -> isPR 2 (fun x y : nat => g (f1 x y) (f2 x y) (f3 x y) (f4 x y)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+induction H1 as (x1, p1); simpl in p1.
+induction H2 as (x2, p2); simpl in p2.
+induction H3 as (x3, p3); simpl in p3.
+exists
+ (composeFunc _ _
+ (PRcons _ _ x (PRcons _ _ x0 (PRcons _ _ x1 (PRcons _ _ x2 (PRnil _)))))
+ x3).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+rewrite <- p1.
+rewrite <- p2.
+rewrite <- p3.
+auto.
+Qed.
+
+Lemma compose3_1IsPR :
+ forall f : nat -> nat -> nat -> nat,
+ isPR 3 f ->
+ forall g : nat -> nat, isPR 1 g -> isPR 3 (fun x y z : nat => g (f x y z)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+exists (composeFunc _ _ (PRcons _ _ x (PRnil _)) x0).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+auto.
+Qed.
+
+Lemma compose3_2IsPR :
+ forall f1 : nat -> nat -> nat -> nat,
+ isPR 3 f1 ->
+ forall f2 : nat -> nat -> nat -> nat,
+ isPR 3 f2 ->
+ forall g : nat -> nat -> nat,
+ isPR 2 g -> isPR 3 (fun x y z : nat => g (f1 x y z) (f2 x y z)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+induction H1 as (x1, p1); simpl in p1.
+exists (composeFunc _ _ (PRcons _ _ x (PRcons _ _ x0 (PRnil _))) x1).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+rewrite <- p1.
+auto.
+Qed.
+
+Lemma compose3_3IsPR :
+ forall f1 : nat -> nat -> nat -> nat,
+ isPR 3 f1 ->
+ forall f2 : nat -> nat -> nat -> nat,
+ isPR 3 f2 ->
+ forall f3 : nat -> nat -> nat -> nat,
+ isPR 3 f3 ->
+ forall g : nat -> nat -> nat -> nat,
+ isPR 3 g -> isPR 3 (fun x y z : nat => g (f1 x y z) (f2 x y z) (f3 x y z)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+induction H1 as (x1, p1); simpl in p1.
+induction H2 as (x2, p2); simpl in p2.
+exists
+ (composeFunc _ _ (PRcons _ _ x (PRcons _ _ x0 (PRcons _ _ x1 (PRnil _)))) x2).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+rewrite <- p1.
+rewrite <- p2.
+auto.
+Qed.
+
+Lemma compose4_2IsPR :
+ forall f1 : nat -> nat -> nat -> nat -> nat,
+ isPR 4 f1 ->
+ forall f2 : nat -> nat -> nat -> nat -> nat,
+ isPR 4 f2 ->
+ forall g : nat -> nat -> nat,
+ isPR 2 g -> isPR 4 (fun w x y z : nat => g (f1 w x y z) (f2 w x y z)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+induction H1 as (x1, p1); simpl in p1.
+exists (composeFunc _ _ (PRcons _ _ x (PRcons _ _ x0 (PRnil _))) x1).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+rewrite <- p1.
+auto.
+Qed.
+
+Lemma compose4_3IsPR :
+ forall f1 : nat -> nat -> nat -> nat -> nat,
+ isPR 4 f1 ->
+ forall f2 : nat -> nat -> nat -> nat -> nat,
+ isPR 4 f2 ->
+ forall f3 : nat -> nat -> nat -> nat -> nat,
+ isPR 4 f3 ->
+ forall g : nat -> nat -> nat -> nat,
+ isPR 3 g ->
+ isPR 4 (fun w x y z : nat => g (f1 w x y z) (f2 w x y z) (f3 w x y z)).
+Proof.
+intros.
+induction H as (x, p); simpl in p.
+induction H0 as (x0, p0); simpl in p0.
+induction H1 as (x1, p1); simpl in p1.
+induction H2 as (x2, p2); simpl in p2.
+exists
+ (composeFunc _ _ (PRcons _ _ x (PRcons _ _ x0 (PRcons _ _ x1 (PRnil _)))) x2).
+simpl in |- *.
+intros.
+rewrite <- p.
+rewrite <- p0.
+rewrite <- p1.
+rewrite <- p2.
+auto.
+Qed.
+
+Lemma swapIsPR :
+ forall f : nat -> nat -> nat, isPR 2 f -> isPR 2 (fun x y : nat => f y x).
+Proof.
+intros.
+apply compose2_2IsPR with (f := fun a b : nat => b) (g := fun a b : nat => a).
+apply pi2_2IsPR.
+apply pi1_2IsPR.
+apply H.
+Qed.
+
+Lemma indIsPR :
+ forall f : nat -> nat -> nat,
+ isPR 2 f ->
+ forall g : nat,
+ isPR 1
+ (fun a : nat => nat_rec (fun n : nat => nat) g (fun x y : nat => f x y) a).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+induction (const0_NIsPR g).
+simpl in p0.
+exists (primRecFunc _ x0 x).
+simpl in |- *.
+simple induction c.
+simpl in |- *.
+rewrite <- p0.
+auto.
+intros.
+simpl in |- *.
+rewrite <- p.
+rewrite <- H.
+auto.
+Qed.
+
+Lemma ind1ParamIsPR :
+ forall f : nat -> nat -> nat -> nat,
+ isPR 3 f ->
+ forall g : nat -> nat,
+ isPR 1 g ->
+ isPR 2
+ (fun a b : nat =>
+ nat_rec (fun n : nat => nat) (g b) (fun x y : nat => f x y b) a).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+induction H0 as (x0, p0).
+simpl in p0.
+exists (primRecFunc _ x0 x).
+simpl in |- *.
+intros.
+induction c as [| c Hrecc].
+simpl in |- *.
+apply p0.
+simpl in |- *.
+rewrite p.
+rewrite Hrecc.
+auto.
+Qed.
+
+Lemma ind2ParamIsPR :
+ forall f : nat -> nat -> nat -> nat -> nat,
+ isPR 4 f ->
+ forall g : nat -> nat -> nat,
+ isPR 2 g ->
+ isPR 3
+ (fun a b c : nat =>
+ nat_rec (fun n : nat => nat) (g b c) (fun x y : nat => f x y b c) a).
+Proof.
+intros.
+induction H as (x, p).
+simpl in p.
+induction H0 as (x0, p0).
+simpl in p0.
+exists (primRecFunc _ x0 x).
+simpl in |- *.
+simple induction c.
+intros.
+simpl in |- *.
+rewrite p0.
+auto.
+intros.
+simpl in |- *.
+rewrite p.
+rewrite H.
+auto.
+Qed.
+
+Lemma plusIndIsPR : isPR 3 (fun n fn b : nat => S fn).
+apply (filter010IsPR _ succIsPR).
+Qed.
+
+Lemma plusIsPR : isPR 2 plus.
+Proof.
+assert
+ (isPR 2
+ (fun a b : nat => nat_rec (fun n : nat => nat) b (fun x y : nat => S y) a)).
+apply (ind1ParamIsPR _ plusIndIsPR _ idIsPR).
+induction H as (x, p).
+simpl in p.
+exists x.
+simpl in |- *.
+intros.
+rewrite p.
+induction c as [| c Hrecc].
+auto.
+simpl in |- *.
+rewrite Hrecc.
+auto.
+Qed.
+
+Lemma multIndIsPR : isPR 3 (fun n fn b : nat => fn + b).
+apply (filter011IsPR _ plusIsPR).
+Qed.
+
+Lemma multIsPR : isPR 2 mult.
+Proof.
+assert
+ (isPR 2
+ (fun a b : nat =>
+ nat_rec (fun n : nat => nat) 0 (fun x y : nat => y + b) a)).
+assert (isPR 1 (fun _ => 0)).
+apply const1_NIsPR.
+apply (ind1ParamIsPR _ multIndIsPR _ H).
+induction H as (x, p).
+simpl in p.
+exists x.
+simpl in |- *.
+intros.
+rewrite p.
+induction c as [| c Hrecc].
+auto.
+simpl in |- *.
+rewrite Hrecc.
+apply plus_comm.
+Qed.
+
+Lemma predIsPR : isPR 1 pred.
+Proof.
+assert
+ (isPR 1
+ (fun a : nat => nat_rec (fun n : nat => nat) 0 (fun x y : nat => x) a)).
+apply (indIsPR _ pi1_2IsPR 0).
+induction H as (x, p).
+simpl in p.
+exists x.
+simpl in |- *.
+intros.
+rewrite p.
+induction c as [| c Hrecc].
+auto.
+auto.
+Qed.
+
+Lemma minusIndIsPR : isPR 3 (fun n fn b : nat => pred fn).
+apply (filter010IsPR _ predIsPR).
+Qed.
+
+Lemma minusIsPR : isPR 2 minus.
+Proof.
+assert
+ (isPR 2
+ (fun b a : nat =>
+ nat_rec (fun n : nat => nat) b (fun x y : nat => pred y) a)).
+apply
+ swapIsPR
+ with
+ (f := fun a b : nat =>
+ nat_rec (fun n : nat => nat) b (fun x y : nat => pred y) a).
+apply (ind1ParamIsPR _ minusIndIsPR _ idIsPR).
+induction H as (x, p).
+simpl in p.
+exists x.
+simpl in |- *.
+intros.
+rewrite p.
+induction c0 as [| c0 Hrecc0].
+simpl in |- *.
+apply minus_n_O.
+simpl in |- *.
+rewrite Hrecc0.
+generalize c c0.
+intro.
+induction c1 as [| c1 Hrecc1].
+intros.
+auto.
+intros.
+simpl in |- *.
+induction c2 as [| c2 Hrecc2].
+simpl in |- *.
+apply minus_n_O.
+apply Hrecc1.
+Qed.
+
+Definition notZero (a : nat) :=
+ nat_rec (fun n : nat => nat) 0 (fun x y : nat => 1) a.
+
+Lemma notZeroIsPR : isPR 1 notZero.
+Proof.
+unfold notZero in |- *.
+apply indIsPR with (f := fun _ _ : nat => 1).
+apply filter10IsPR with (g := fun _ : nat => 1).
+apply const1_NIsPR.
+Qed.
+
+Definition ltBool (a b : nat) : bool.
+intros.
+destruct (le_lt_dec b a).
+exact false.
+exact true.
+Defined.
+
+Lemma ltBoolTrue : forall a b : nat, ltBool a b = true -> a < b.
+intros.
+unfold ltBool in H.
+induction (le_lt_dec b a).
+discriminate H.
+auto.
+Qed.
+
+Lemma ltBoolFalse : forall a b : nat, ltBool a b = false -> ~ a < b.
+intros.
+unfold ltBool in H.
+induction (le_lt_dec b a).
+unfold not in |- *; intros.
+elim (le_not_lt _ _ a0).
+auto.
+discriminate H.
+Qed.
+
+Lemma ltIsPR : isPRrel 2 ltBool.
+Proof.
+unfold isPRrel in |- *.
+assert (isPR 2 (fun a b : nat => notZero (b - a))).
+apply swapIsPR with (f := fun a b : nat => notZero (a - b)).
+apply compose2_1IsPR.
+apply minusIsPR.
+apply notZeroIsPR.
+induction H as (x, p).
+simpl in p.
+exists x.
+simpl in |- *.
+intros.
+rewrite p.
+unfold ltBool in |- *.
+induction (le_lt_dec c0 c).
+cut (c0 <= c).
+generalize c.
+clear c a.
+induction c0 as [| c0 Hrecc0].
+intros.
+reflexivity.
+intros.
+induction c as [| c Hrecc].
+elim (le_Sn_O _ H).
+simpl in |- *.
+apply Hrecc0.
+apply le_S_n.
+auto.
+auto.
+cut (c < c0).
+generalize c.
+clear c b.
+induction c0 as [| c0 Hrecc0].
+intros.
+elim (lt_n_O _ H).
+intros.
+induction c as [| c Hrecc].
+simpl in |- *.
+reflexivity.
+simpl in |- *.
+apply Hrecc0.
+apply lt_S_n.
+auto.
+auto.
+Qed.
+
+Lemma maxIsPR : isPR 2 max.
+Proof.
+assert (isPR 2 (fun a b : nat => a + (b - a))).
+apply
+ compose2_2IsPR
+ with (h := plus) (f := fun a b : nat => a) (g := fun a b : nat => b - a).
+apply pi1_2IsPR.
+apply swapIsPR.
+apply minusIsPR.
+apply plusIsPR.
+induction H as (x, p).
+exists x.
+eapply extEqualTrans.
+apply p.
+clear x p.
+simpl in |- *.
+intros.
+induction (le_or_lt c c0).
+rewrite max_r.
+symmetry in |- *.
+apply le_plus_minus.
+assumption.
+assumption.
+rewrite not_le_minus_0.
+rewrite plus_comm.
+rewrite max_l.
+reflexivity.
+apply lt_le_weak.
+assumption.
+apply lt_not_le.
+assumption.
+Qed.
+
+Lemma gtIsPR : isPRrel 2 (fun a b : nat => ltBool b a).
+Proof.
+intros.
+unfold isPRrel in |- *.
+simpl in |- *.
+apply swapIsPR with (f := fun a0 a : nat => if ltBool a0 a then 1 else 0).
+apply ltIsPR.
+Qed.
+
+Remark replaceCompose2 :
+ forall (n : nat) (a b a' b' : naryFunc n) (c c' : naryFunc 2),
+ extEqual n a a' ->
+ extEqual n b b' ->
+ extEqual 2 c c' ->
+ extEqual n
+ (evalComposeFunc _ _ (Vector.cons _ a _ (Vector.cons _ b _ (Vector.nil (naryFunc n)))) c)
+ (evalComposeFunc _ _ (Vector.cons _ a' _ (Vector.cons _ b' _ (Vector.nil (naryFunc n)))) c').
+Proof.
+intros.
+apply extEqualCompose.
+unfold extEqualVector in |- *.
+simpl in |- *.
+auto.
+auto.
+Qed.
+
+Definition orRel (n : nat) (R S : naryRel n) : naryRel n.
+intros.
+induction n as [| n Hrecn].
+apply (R || S).
+simpl in |- *.
+intros.
+apply Hrecn.
+apply (R H).
+apply (S H).
+Defined.
+
+Lemma orRelPR :
+ forall (n : nat) (R R' : naryRel n),
+ isPRrel n R -> isPRrel n R' -> isPRrel n (orRel n R R').
+Proof.
+intros.
+induction H as (x, p).
+induction H0 as (x0, p0).
+assert (isPR 2 (fun a b : nat => notZero (a + b))).
+apply compose2_1IsPR.
+apply plusIsPR.
+apply notZeroIsPR.
+induction H as (x1, p1).
+exists (composeFunc _ _ (PRcons _ _ x (PRcons _ _ x0 (PRnil _))) x1).
+simpl in |- *.
+apply
+ extEqualTrans
+ with
+ (evalComposeFunc n 2
+ (Vector.cons (naryFunc n) (charFunction n R) 1
+ (Vector.cons (naryFunc n) (charFunction n R') 0 (Vector.nil (naryFunc n))))
+ (fun a b : nat => notZero (a + b))).
+apply replaceCompose2; auto.
+clear x p x0 p0 x1 p1.
+induction n as [| n Hrecn].
+simpl in |- *.
+induction R.
+reflexivity.
+induction R'.
+reflexivity.
+reflexivity.
+simpl in |- *.
+fold (naryFunc n) in |- *.
+intros.
+apply (Hrecn (R c) (R' c)).
+Qed.
+
+Definition andRel (n : nat) (R S : naryRel n) : naryRel n.
+intros.
+induction n as [| n Hrecn].
+exact (R && S).
+simpl in |- *.
+intros.
+apply Hrecn.
+apply (R H).
+apply (S H).
+Defined.
+
+Lemma andRelPR :
+ forall (n : nat) (R R' : naryRel n),
+ isPRrel n R -> isPRrel n R' -> isPRrel n (andRel n R R').
+Proof.
+intros.
+induction H as (x, p).
+induction H0 as (x0, p0).
+assert (isPR 2 mult).
+apply multIsPR.
+induction H as (x1, p1).
+exists (composeFunc _ _ (PRcons _ _ x (PRcons _ _ x0 (PRnil _))) x1).
+simpl in |- *.
+apply
+ extEqualTrans
+ with
+ (evalComposeFunc n 2
+ (Vector.cons (naryFunc n) (charFunction n R) 1
+ (Vector.cons (naryFunc n) (charFunction n R') 0 (Vector.nil (naryFunc n))))
+ mult).
+apply replaceCompose2; auto.
+clear x p x0 p0 x1 p1.
+induction n as [| n Hrecn].
+simpl in |- *.
+induction R.
+induction R'.
+reflexivity.
+reflexivity.
+reflexivity.
+simpl in |- *.
+fold (naryFunc n) in |- *.
+intros.
+apply (Hrecn (R c) (R' c)).
+Qed.
+
+Definition notRel (n : nat) (R : naryRel n) : naryRel n.
+intros.
+induction n as [| n Hrecn].
+exact (negb R).
+simpl in |- *.
+intros.
+apply Hrecn.
+apply (R H).
+Defined.
+
+Lemma notRelPR :
+ forall (n : nat) (R : naryRel n), isPRrel n R -> isPRrel n (notRel n R).
+Proof.
+intros.
+induction H as (x, p).
+assert (isPR 1 (fun x : nat => 1 - x)).
+apply compose1_2IsPR with (f := fun x : nat => 1) (f' := fun x : nat => x).
+apply const1_NIsPR.
+apply idIsPR.
+apply minusIsPR.
+induction H as (x0, p0).
+exists (composeFunc _ _ (PRcons _ _ x (PRnil _)) x0).
+simpl in |- *.
+apply
+ extEqualTrans
+ with
+ (evalComposeFunc n 1 (Vector.cons _ (charFunction n R) _ (Vector.nil _))
+ (fun x : nat => 1 - x)).
+apply extEqualCompose.
+unfold extEqualVector in |- *.
+simpl in |- *.
+auto.
+auto.
+clear p0 x0 p x.
+induction n as [| n Hrecn].
+simpl in |- *.
+induction R.
+reflexivity.
+reflexivity.
+simpl in |- *.
+intros.
+fold (naryFunc n) in |- *.
+apply Hrecn.
+Qed.
+
+Fixpoint bodd (n : nat) : bool :=
+ match n with
+ | O => false
+ | S n' => negb (bodd n')
+ end.
+
+Lemma boddIsPR : isPRrel 1 bodd.
+Proof.
+assert (isPR 2 (fun _ rec : nat => 1 - rec)).
+apply filter01IsPR with (g := fun rec : nat => 1 - rec).
+apply compose1_2IsPR with (f := fun _ : nat => 1) (f' := fun x : nat => x).
+apply const1_NIsPR.
+apply idIsPR.
+apply minusIsPR.
+induction H as (x, p).
+exists (primRecFunc 0 zeroFunc x).
+simpl in |- *.
+intros.
+induction c as [| c Hrecc].
+simpl in |- *.
+auto.
+simpl in |- *.
+rewrite p.
+simpl in |- *.
+rewrite Hrecc.
+clear Hrecc.
+induction (bodd c).
+reflexivity.
+reflexivity.
+Qed.
+
+Lemma beq_nat_not_refl : forall a b : nat, a <> b -> beq_nat a b = false.
+Proof.
+double induction a b.
+intros.
+elim H.
+auto.
+auto.
+auto.
+intros.
+simpl in |- *.
+auto.
+Qed.
+
+Lemma neqIsPR : isPRrel 2 (fun a b : nat => negb (beq_nat a b)).
+Proof.
+intros.
+assert (isPRrel 2 (orRel 2 ltBool (fun a b : nat => ltBool b a))).
+apply orRelPR.
+apply ltIsPR.
+apply gtIsPR.
+induction H as (x, p).
+exists x.
+simpl in |- *.
+simpl in p.
+intros.
+rewrite p.
+clear p.
+unfold ltBool in |- *.
+induction (eq_nat_dec c c0).
+rewrite a.
+rewrite <- beq_nat_refl.
+simpl in |- *.
+induction (le_lt_dec c0 c0).
+simpl in |- *.
+reflexivity.
+elim (lt_irrefl _ b).
+rewrite beq_nat_not_refl.
+simpl in |- *.
+induction (le_lt_dec c0 c).
+induction (le_lt_dec c c0).
+induction (nat_total_order _ _ b).
+elim (lt_not_le _ _ H).
+auto.
+elim (lt_not_le _ _ H).
+auto.
+reflexivity.
+reflexivity.
+auto.
+Qed.
+
+Lemma eqIsPR : isPRrel 2 beq_nat.
+Proof.
+intros.
+assert (isPRrel 2 (notRel 2 (fun a b : nat => negb (beq_nat a b)))).
+apply notRelPR.
+apply neqIsPR.
+simpl in H.
+induction H as (x, p).
+exists x.
+simpl in |- *.
+simpl in p.
+intros.
+rewrite p.
+clear p.
+induction (beq_nat c c0).
+auto.
+auto.
+Qed.
+
+Definition leBool (a b : nat) : bool.
+intros.
+destruct (le_lt_dec a b).
+exact true.
+exact false.
+Defined.
+
+Lemma leIsPR : isPRrel 2 leBool.
+Proof.
+assert (isPRrel 2 (orRel 2 ltBool beq_nat)).
+apply orRelPR.
+apply ltIsPR.
+apply eqIsPR.
+induction H as (x, p).
+exists x.
+simpl in |- *.
+simpl in p.
+intros.
+rewrite p.
+clear p x.
+unfold leBool in |- *.
+unfold ltBool in |- *.
+induction (le_lt_dec c0 c).
+induction (le_lt_dec c c0).
+simpl in |- *.
+replace c0 with c.
+rewrite <- beq_nat_refl.
+reflexivity.
+induction (eq_nat_dec c c0).
+auto.
+induction (nat_total_order _ _ b).
+elim (lt_not_le _ _ H).
+auto.
+elim (lt_not_le _ _ H).
+auto.
+rewrite beq_nat_not_refl.
+simpl in |- *.
+reflexivity.
+unfold not in |- *; intros.
+rewrite H in b.
+elim (lt_irrefl _ b).
+simpl in |- *.
+induction (le_lt_dec c c0).
+reflexivity.
+elim (lt_irrefl c).
+apply lt_trans with c0; auto.
+Qed.
+
+Section Ignore_Params.
+
+Fixpoint ignoreParams (n m : nat) (f : naryFunc n) {struct m} :
+ naryFunc (m + n) :=
+ match m return (naryFunc (m + n)) with
+ | O => f
+ | S x => fun _ => ignoreParams n x f
+ end.
+
+Definition projectionListPR (n m : nat) (p : m <= n) : PrimRecs n m.
+intros.
+induction m as [| m Hrecm].
+exact (PRnil n).
+assert (m < n).
+apply lt_S_n.
+apply le_lt_n_Sm.
+apply p.
+apply (PRcons _ m (projFunc _ _ H)).
+apply Hrecm.
+apply le_S_n.
+apply le_S.
+apply p.
+Defined.
+
+Definition projectionList (n m : nat) (p : m <= n) :
+ Vector.t (naryFunc n) m := evalPrimRecs n m (projectionListPR n m p).
+
+Lemma projectionListInd :
+ forall (n m : nat) (p1 p2 : m <= n),
+ projectionList n m p1 = projectionList n m p2.
+Proof.
+intros.
+unfold projectionList in |- *.
+induction m as [| m Hrecm].
+simpl in |- *.
+reflexivity.
+simpl in |- *.
+rewrite (Hrecm (le_S_n m n (le_S (S m) n p1)) (le_S_n m n (le_S (S m) n p2))).
+rewrite
+ (evalProjFuncInd _ _ (lt_S_n m n (le_lt_n_Sm (S m) n p1))
+ (lt_S_n m n (le_lt_n_Sm (S m) n p2))).
+reflexivity.
+Qed.
+
+Lemma projectionListApplyParam :
+ forall (m n c : nat) (p1 : m <= n) (p2 : m <= S n),
+ extEqualVector _ _ (projectionList n m p1)
+ (evalOneParamList n m c (projectionList (S n) m p2)).
+Proof.
+unfold extEqualVector in |- *.
+unfold projectionList in |- *.
+simple induction m.
+simpl in |- *.
+auto.
+intros.
+simpl in |- *.
+induction (eq_nat_dec n n0).
+elim (lt_not_le n (S n)).
+apply lt_n_Sn.
+rewrite <- a in p1.
+auto.
+split.
+rewrite
+ (evalProjFuncInd _ _ (lt_S_n n n0 (le_lt_n_Sm (S n) n0 p1))
+ match
+ le_lt_or_eq n n0
+ (lt_n_Sm_le n n0 (lt_S_n n (S n0) (le_lt_n_Sm (S n) (S n0) p2)))
+ with
+ | or_introl l2 => l2
+ | or_intror l2 => False_ind (n < n0) (b l2)
+ end).
+apply extEqualRefl.
+apply H.
+Qed.
+
+Lemma projectionListId :
+ forall (n : nat) (f : naryFunc n) (p : n <= n),
+ extEqual n f (evalComposeFunc n n (projectionList n n p) f).
+Proof.
+intros.
+induction n as [| n Hrecn].
+simpl in |- *.
+reflexivity.
+simpl in |- *.
+intros.
+fold (naryFunc n) in |- *.
+induction (eq_nat_dec n n).
+apply
+ extEqualTrans
+ with
+ (evalComposeFunc n (S n)
+ (Vector.cons (naryFunc n) (evalConstFunc n c) n
+ (projectionList n n (le_n n))) f).
+apply
+ extEqualTrans with (evalComposeFunc n n (projectionList n n (le_n n)) (f c)).
+apply Hrecn.
+clear Hrecn a p.
+generalize (projectionList n n (le_n n)).
+generalize f c.
+clear f c.
+assert
+ (forall (m : nat) (f : naryFunc (S m)) (c : nat) (v : Vector.t (naryFunc n) m),
+ extEqual n (evalComposeFunc n m v (f c))
+ (evalComposeFunc n (S m) (Vector.cons (naryFunc n) (evalConstFunc n c) m v) f)).
+intros.
+induction n as [| n Hrecn].
+simpl in |- *.
+reflexivity.
+simpl in |- *.
+intros.
+fold (naryFunc n) in |- *.
+apply Hrecn.
+apply H.
+apply extEqualCompose.
+unfold extEqualVector in |- *.
+simpl in |- *.
+split.
+apply extEqualRefl.
+apply
+ (projectionListApplyParam n n c (le_n n)
+ (le_S_n n (S n) (le_S (S n) (S n) p))).
+apply extEqualRefl.
+elim b.
+auto.
+Qed.
+
+Lemma ignoreParamsIsPR :
+ forall (n m : nat) (f : naryFunc n), isPR _ f -> isPR _ (ignoreParams n m f).
+Proof.
+assert
+ (forall (n m : nat) (pr : m <= n) (f : naryFunc m) (p : n - m + m = n),
+ extEqual _ (eq_rec (n - m + m) naryFunc (ignoreParams m (n - m) f) _ p)
+ (evalComposeFunc _ _ (projectionList n m pr) f)).
+unfold projectionList in |- *.
+intro.
+induction n as [| n Hrecn]; intros.
+destruct m as [| n].
+simpl in |- *.
+elim p using K_dec_set.
+apply eq_nat_dec.
+simpl in |- *.
+reflexivity.
+elim (le_Sn_O _ pr).
+induction (le_lt_or_eq _ _ pr).
+assert (m <= n).
+apply lt_n_Sm_le.
+auto.
+generalize p.
+rewrite <- minus_Sn_m.
+clear p.
+intros.
+simpl in |- *.
+intros.
+assert (n - m + m = n).
+simpl in p.
+injection p.
+auto.
+replace
+ (eq_rec (S (n - m + m)) naryFunc (fun _ : nat => ignoreParams m (n - m) f)
+ (S n) p c) with
+ (eq_rec (n - m + m) naryFunc (ignoreParams m (n - m) f) n H1).
+apply extEqualTrans with (evalComposeFunc n m (projectionList n m H0) f).
+unfold projectionList in |- *.
+apply Hrecn.
+apply extEqualCompose.
+apply (projectionListApplyParam m n c H0 pr).
+apply extEqualRefl.
+generalize H1.
+generalize p.
+simpl in |- *.
+generalize (ignoreParams m (n - m) f).
+rewrite H1.
+intros.
+elim H2 using K_dec_set.
+apply eq_nat_dec.
+elim p0 using K_dec_set.
+apply eq_nat_dec.
+simpl in |- *.
+reflexivity.
+auto.
+generalize p.
+generalize pr.
+rewrite <- H.
+clear H p.
+replace (m - m) with 0.
+simpl in |- *.
+intros.
+elim p using K_dec_set.
+apply eq_nat_dec.
+simpl in |- *.
+clear p pr.
+apply (projectionListId m f pr0).
+apply minus_n_n.
+intros.
+unfold projectionList in H.
+induction H0 as (x, p).
+exists (composeFunc (m + n) n (projectionListPR _ _ (le_plus_r _ _)) x).
+apply extEqualSym.
+assert (m + n - n + n = m + n).
+rewrite (plus_comm m n).
+rewrite minus_plus.
+apply plus_comm.
+assert
+ (extEqual (m + n)
+ (eq_rec (m + n - n + n) naryFunc (ignoreParams n (m + n - n) f)
+ (m + n) H0)
+ (evalComposeFunc (m + n) _
+ (evalPrimRecs (m + n) n (projectionListPR (m + n) n (le_plus_r m n)))
+ f)).
+apply H.
+replace (ignoreParams n m f) with
+ (eq_rec (m + n - n + n) naryFunc (ignoreParams n (m + n - n) f) (m + n) H0).
+simpl in |- *.
+apply
+ extEqualTrans
+ with
+ (evalComposeFunc (m + n) _
+ (evalPrimRecs (m + n) n (projectionListPR (m + n) n (le_plus_r m n)))
+ f).
+apply H1.
+apply extEqualCompose.
+generalize
+ (evalPrimRecs (m + n) n (projectionListPR (m + n) n (le_plus_r m n))).
+generalize (m + n).
+intros.
+apply extEqualVectorRefl.
+apply extEqualSym.
+auto.
+generalize H0.
+replace (m + n - n) with m.
+intros.
+elim H2 using K_dec_set.
+apply eq_nat_dec.
+simpl in |- *.
+reflexivity.
+rewrite plus_comm.
+rewrite minus_plus.
+reflexivity.
+Qed.
+
+End Ignore_Params.
+
+Lemma reduce1stCompose :
+ forall (c n m : nat) (v : Vector.t (naryFunc n) m) (g : naryFunc (S m)),
+ extEqual n
+ (evalComposeFunc n _ (Vector.cons (naryFunc n) (evalConstFunc n c) _ v) g)
+ (evalComposeFunc n _ v (g c)).
+Proof.
+intros c n.
+induction n as [| n Hrecn].
+simpl in |- *.
+auto.
+simpl in |- *.
+fold (naryFunc n) in |- *.
+intros.
+apply Hrecn.
+Qed.
+
+Lemma reduce2ndCompose :
+ forall (c n m : nat) (v : Vector.t (naryFunc n) m) (n0 : naryFunc n)
+ (g : naryFunc (S (S m))),
+ extEqual n
+ (evalComposeFunc n _
+ (Vector.cons (naryFunc n) n0 _ (Vector.cons (naryFunc n) (evalConstFunc n c) _ v))
+ g)
+ (evalComposeFunc n _ (Vector.cons (naryFunc n) n0 _ v) (fun x : nat => g x c)).
+Proof.
+intros c n.
+induction n as [| n Hrecn].
+simpl in |- *.
+auto.
+simpl in |- *.
+fold (naryFunc n) in |- *.
+intros.
+apply Hrecn.
+Qed.
+
+Lemma evalPrimRecParam :
+ forall (n c d : nat) (g : naryFunc (S n)) (h : naryFunc (S (S (S n)))),
+ extEqual _ (evalPrimRecFunc n (g d) (fun x y : nat => h x y d) c)
+ (evalPrimRecFunc (S n) g h c d).
+Proof.
+intros.
+induction c as [| c Hrecc].
+simpl in |- *.
+apply extEqualRefl.
+simpl in |- *.
+apply extEqualCompose2.
+auto.
+apply extEqualRefl.
+Qed.
+
+Lemma compose2IsPR :
+ forall (n : nat) (f : naryFunc n) (p : isPR n f) (g : naryFunc (S n))
+ (q : isPR (S n) g), isPR n (compose2 n f g).
+Proof.
+intros.
+induction p as (x, p).
+induction q as (x0, p0).
+exists (composeFunc _ _ (PRcons _ _ x (projectionListPR n n (le_n n))) x0).
+simpl in |- *.
+apply
+ extEqualTrans
+ with
+ (evalComposeFunc n (S n)
+ (Vector.cons (naryFunc n) f n
+ (evalPrimRecs n n (projectionListPR n n (le_n n)))) g).
+apply extEqualCompose.
+unfold extEqualVector in |- *.
+simpl in |- *.
+split.
+auto.
+apply extEqualVectorRefl.
+auto.
+clear p0 x0 p x.
+induction n as [| n Hrecn].
+simpl in |- *.
+auto.
+simpl in |- *.
+fold (naryFunc n) in |- *.
+induction (eq_nat_dec n n).
+intros.
+apply
+ extEqualTrans
+ with
+ (evalComposeFunc n (S (S n))
+ (Vector.cons (naryFunc n) (f c) (S n)
+ (Vector.cons (naryFunc n) (evalConstFunc n c) n
+ (projectionList n n (le_n n)))) g).
+apply extEqualSym.
+apply extEqualCompose.
+unfold extEqualVector in |- *.
+simpl in |- *.
+repeat split.
+apply extEqualRefl.
+apply extEqualRefl.
+apply
+ (projectionListApplyParam n n c (le_n n)
+ (le_S_n n (S n) (le_S (S n) (S n) (le_n (S n))))).
+apply extEqualRefl.
+apply
+ extEqualTrans
+ with
+ (evalComposeFunc n (S n)
+ (Vector.cons (naryFunc n) (f c) n
+ (evalPrimRecs n n (projectionListPR n n (le_n n))))
+ (fun x : nat => g x c)).
+clear a Hrecn.
+generalize (f c).
+fold (naryFunc n) in |- *.
+fold (projectionList n n (le_n n)) in |- *.
+generalize (projectionList n n (le_n n)).
+intros v n0.
+apply (reduce2ndCompose c n n v n0).
+apply Hrecn.
+elim b.
+auto.
+Qed.
+
+Lemma compose1_NIsPR :
+ forall (n : nat) (g : naryFunc (S n)),
+ isPR (S n) g ->
+ forall f : naryFunc 1, isPR 1 f -> isPR (S n) (fun x : nat => g (f x)).
+Proof.
+intros.
+induction H as (x, p).
+induction H0 as (x0, p0).
+exists
+ (composeFunc (S n) (S n)
+ (PRcons _ _
+ (composeFunc (S n) 1
+ (PRcons _ _ (projFunc (S n) n (lt_n_Sn n)) (PRnil _)) x0)
+ (projectionListPR (S n) n (le_n_Sn n))) x).
+simpl in |- *.
+fold (naryFunc n) in |- *.
+induction (eq_nat_dec n n).
+intros.
+apply
+ extEqualTrans
+ with
+ (evalComposeFunc n (S n)
+ (Vector.cons (naryFunc n) (evalConstFunc n (f c)) n
+ (projectionList n n (le_n n))) g).
+apply extEqualSym.
+apply extEqualCompose.
+unfold extEqualVector in |- *.
+simpl in |- *.
+split.
+apply extEqualSym.
+apply
+ extEqualTrans
+ with
+ (evalComposeFunc n 1
+ (Vector.cons (naryFunc n) (evalConstFunc n c) 0 (Vector.nil (naryFunc n))) f).
+apply extEqualCompose.
+apply extEqualVectorRefl.
+auto.
+clear a p0 x0 p x g.
+induction n as [| n Hrecn].
+simpl in |- *.
+reflexivity.
+simpl in |- *.
+intros.
+apply Hrecn.
+apply (projectionListApplyParam n n c (le_n n) (le_n_Sn n)).
+apply extEqualSym.
+auto.
+clear p0 x0 p x a.
+eapply extEqualTrans.
+apply reduce1stCompose.
+apply extEqualSym.
+apply (projectionListId n (g (f c)) (le_n n)).
+elim b.
+auto.
+Qed.
+
+Definition switchPR : naryFunc 3.
+simpl in |- *.
+intros.
+destruct H as [| n].
+apply H1.
+apply H0.
+Defined.
+
+Lemma switchIsPR : isPR 3 switchPR.
+Proof.
+intros.
+assert
+ (isPR 3
+ (fun a b c : nat =>
+ nat_rec (fun _ : nat => nat) c
+ (fun (n : nat) (_ : (fun _ : nat => nat) n) => b) a)).
+apply
+ ind2ParamIsPR with (f := fun _ _ b c : nat => b) (g := fun b c : nat => c).
+apply pi3_4IsPR.
+apply pi2_2IsPR.
+induction H as (x, p).
+exists x.
+apply
+ extEqualTrans
+ with
+ (fun a b c : nat => nat_rec (fun _ : nat => nat) c (fun _ _ : nat => b) a).
+apply p.
+unfold switchPR in |- *.
+simpl in |- *.
+intros.
+induction c; reflexivity.
+Qed.
+
+(* Returns smallest value of x below b such that (P b x). Otherwise returns b*)
+Fixpoint boundedSearchHelp (P : naryRel 1) (b : nat) {struct b} : nat :=
+ match b with
+ | O => 0
+ | S b' =>
+ match eq_nat_dec (boundedSearchHelp P b') b' with
+ | left _ => match P b' with
+ | true => b'
+ | false => S b'
+ end
+ | right _ => boundedSearchHelp P b'
+ end
+ end.
+
+Definition boundedSearch (P : naryRel 2) (b : nat) : nat :=
+ boundedSearchHelp (P b) b.
+
+Lemma boundedSearch1 :
+ forall (P : naryRel 2) (b x : nat), x < boundedSearch P b -> P b x = false.
+Proof.
+unfold boundedSearch in |- *.
+intros P b.
+generalize (P b).
+intros b0 x H.
+clear P.
+induction b as [| b Hrecb].
+simpl in H.
+elim (lt_n_O _ H).
+simpl in H.
+induction (eq_nat_dec (boundedSearchHelp b0 b) b).
+rewrite a in Hrecb.
+induction (eq_nat_dec x b).
+rewrite a0.
+induction (b0 b).
+elim (lt_irrefl b).
+rewrite a0 in H.
+auto.
+auto.
+apply Hrecb.
+induction (b0 b).
+auto.
+assert (x <= b).
+apply lt_n_Sm_le.
+auto.
+induction (le_lt_or_eq _ _ H0).
+auto.
+elim b1.
+auto.
+apply Hrecb.
+auto.
+Qed.
+
+Lemma boundedSearch2 :
+ forall (P : naryRel 2) (b : nat),
+ boundedSearch P b = b \/ P b (boundedSearch P b) = true.
+Proof.
+unfold boundedSearch in |- *.
+intros P b.
+assert
+ (forall P : naryRel 1,
+ boundedSearchHelp P b = b \/ P (boundedSearchHelp P b) = true).
+clear P.
+intros.
+induction b as [| b Hrecb].
+simpl in |- *.
+auto.
+simpl in |- *.
+assert (P b = true \/ P b = false).
+induction (P b); auto.
+induction (eq_nat_dec (boundedSearchHelp P b) b).
+induction H as [H| H].
+rewrite H.
+rewrite H.
+auto.
+rewrite H.
+auto.
+induction Hrecb as [H0| H0].
+elim b0.
+auto.
+auto.
+apply H.
+Qed.
+
+Lemma boundSearchIsPR :
+ forall P : naryRel 2,
+ isPRrel 2 P -> isPR 1 (fun b : nat => boundedSearch P b).
+Proof.
+intros.
+unfold boundedSearch in |- *.
+assert
+ (isPR 2
+ (fun b c : nat =>
+ nat_rec (fun _ : nat => nat) 0
+ (fun b0 Hrecb : nat =>
+ sumbool_rec (fun _ : {Hrecb = b0} + {Hrecb <> b0} => nat)
+ (fun _ : Hrecb = b0 =>
+ bool_rec (fun _ : bool => nat) b0 (S b0) (P c b0))
+ (fun _ : Hrecb <> b0 => Hrecb) (eq_nat_dec Hrecb b0)) b)).
+apply
+ ind1ParamIsPR
+ with
+ (g := fun _ : nat => 0)
+ (f := fun b0 Hrecb c : nat =>
+ sumbool_rec (fun _ : {Hrecb = b0} + {Hrecb <> b0} => nat)
+ (fun _ : Hrecb = b0 =>
+ bool_rec (fun _ : bool => nat) b0 (S b0) (P c b0))
+ (fun _ : Hrecb <> b0 => Hrecb) (eq_nat_dec Hrecb b0)).
+unfold isPRrel in H.
+assert
+ (isPR 3
+ (fun b0 Hrecb c : nat =>
+ switchPR (charFunction 2 beq_nat Hrecb b0)
+ (bool_rec (fun _ : bool => nat) b0 (S b0) (P c b0)) Hrecb)).
+apply
+ compose3_3IsPR
+ with
+ (g := switchPR)
+ (f1 := fun b0 Hrecb c : nat => charFunction 2 beq_nat Hrecb b0)
+ (f2 := fun b0 Hrecb c : nat =>
+ bool_rec (fun _ : bool => nat) b0 (S b0) (P c b0))
+ (f3 := fun b0 Hrecb c : nat => Hrecb).
+apply
+ filter110IsPR
+ with (g := fun b0 Hrecb : nat => charFunction 2 beq_nat Hrecb b0).
+apply swapIsPR.
+apply eqIsPR.
+apply
+ filter101IsPR
+ with
+ (g := fun b0 c : nat => bool_rec (fun _ : bool => nat) b0 (S b0) (P c b0)).
+assert
+ (isPR 2 (fun b0 c : nat => switchPR (charFunction 2 P c b0) b0 (S b0))).
+apply
+ compose2_3IsPR
+ with
+ (g := switchPR)
+ (f1 := fun b0 c : nat => charFunction 2 P c b0)
+ (f2 := fun b0 c : nat => b0)
+ (f3 := fun b0 c : nat => S b0).
+apply swapIsPR.
+apply H.
+apply pi1_2IsPR.
+apply filter10IsPR.
+apply succIsPR.
+apply switchIsPR.
+induction H0 as (x, p).
+exists x.
+apply
+ extEqualTrans
+ with (fun b0 c : nat => switchPR (charFunction 2 P c b0) b0 (S b0)).
+auto.
+simpl in |- *.
+intros.
+induction (P c0 c); reflexivity.
+apply pi2_3IsPR.
+apply switchIsPR.
+induction H0 as (x, p).
+exists x.
+apply
+ extEqualTrans
+ with
+ (fun b0 Hrecb c : nat =>
+ switchPR (charFunction 2 beq_nat Hrecb b0)
+ (bool_rec (fun _ : bool => nat) b0 (S b0) (P c b0)) Hrecb).
+auto.
+simpl in |- *.
+intros.
+induction (eq_nat_dec c0 c).
+simpl in |- *.
+rewrite <- a.
+rewrite <- beq_nat_refl.
+simpl in |- *.
+reflexivity.
+rewrite beq_nat_not_refl.
+simpl in |- *.
+reflexivity.
+auto.
+apply const1_NIsPR.
+assert
+ (isPR 1
+ (fun b : nat =>
+ nat_rec (fun _ : nat => nat) 0
+ (fun b0 Hrecb : nat =>
+ sumbool_rec (fun _ : {Hrecb = b0} + {Hrecb <> b0} => nat)
+ (fun _ : Hrecb = b0 =>
+ bool_rec (fun _ : bool => nat) b0 (S b0) (P b b0))
+ (fun _ : Hrecb <> b0 => Hrecb) (eq_nat_dec Hrecb b0)) b)).
+apply
+ compose1_2IsPR
+ with
+ (g := fun b c : nat =>
+ nat_rec (fun _ : nat => nat) 0
+ (fun b0 Hrecb : nat =>
+ sumbool_rec (fun _ : {Hrecb = b0} + {Hrecb <> b0} => nat)
+ (fun _ : Hrecb = b0 =>
+ bool_rec (fun _ : bool => nat) b0 (S b0) (P c b0))
+ (fun _ : Hrecb <> b0 => Hrecb) (eq_nat_dec Hrecb b0)) b)
+ (f := fun b : nat => b)
+ (f' := fun b : nat => b).
+apply idIsPR.
+apply idIsPR.
+auto.
+clear H0.
+induction H1 as (x, p).
+exists x.
+apply
+ extEqualTrans
+ with
+ (fun b : nat =>
+ nat_rec (fun _ : nat => nat) 0
+ (fun b0 Hrecb : nat =>
+ sumbool_rec (fun _ : {Hrecb = b0} + {Hrecb <> b0} => nat)
+ (fun _ : Hrecb = b0 =>
+ bool_rec (fun _ : bool => nat) b0 (S b0) (P b b0))
+ (fun _ : Hrecb <> b0 => Hrecb) (eq_nat_dec Hrecb b0)) b).
+auto.
+clear H x p.
+simpl in |- *.
+intros.
+generalize (P c).
+intros b.
+clear P.
+induction c as [| c Hrecc].
+simpl in |- *.
+auto.
+simpl in |- *.
+rewrite Hrecc.
+induction (eq_nat_dec (boundedSearchHelp b c) c).
+simpl in |- *.
+induction (b c).
+auto.
+auto.
+simpl in |- *.
+reflexivity.
+Qed.
+
+Definition iterate (g : nat -> nat) :=
+ nat_rec (fun _ => nat -> nat) (fun x : nat => x)
+ (fun (_ : nat) (rec : nat -> nat) (x : nat) => g (rec x)).
+
+Lemma iterateIsPR :
+ forall g : nat -> nat, isPR 1 g -> forall n : nat, isPR 1 (iterate g n).
+Proof.
+intros.
+induction n as [| n Hrecn]; simpl in |- *.
+apply idIsPR.
+apply compose1_1IsPR; assumption.
+Qed.
diff --git a/kripke_completeness/Makefile b/kripke_completeness/Makefile
new file mode 100644
index 0000000..83931ce
--- /dev/null
+++ b/kripke_completeness/Makefile
@@ -0,0 +1,20 @@
+
+COQC=coqc
+COQDOC=coqdoc
+DVIPDF=dvipdf
+
+all: main doc
+
+main:
+ $(COQC) stdlib_Type.v int_comp_noquant.v int_comp.v class_comp_cbn_noquant.v class_comp_cbv.v
+
+doc:
+ $(COQDOC) -d html -g --utf8 --toc --no-index *.v
+
+doc-pdf:
+ $(COQDOC) -p "\usepackage{hyperref}\hypersetup{colorlinks=true,linkcolor=red}" -o kripke_completeness.dvi --dvi -toc -g *.v
+ $(DVIPDF) kripke_completeness.dvi
+
+clean:
+ rm -f *.vo *.glob html/class*.html html/int*.html html/std*.html html/coqdoc.css html/toc.html *.dvi *.pdf
+
diff --git a/kripke_completeness/class_comp_cbn_noquant.v b/kripke_completeness/class_comp_cbn_noquant.v
new file mode 100644
index 0000000..aee7fbd
--- /dev/null
+++ b/kripke_completeness/class_comp_cbn_noquant.v
@@ -0,0 +1,1846 @@
+(** Formalised proof of completeness of Classical propositional logic
+ (conjunction, disjunction and implication, no quantifiers) with
+ respect to Kripke-style models.
+
+ Danko Ilik, February 2009
+
+ Normalisation examples at end of file.
+*)
+Require Import stdlib_Type.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+(** printing <= $\le$ #≤# *)
+
+Open Scope type_scope.
+
+Parameters var_trm var_ect : Set.
+
+(** Commands, terms, and evaluation contexts *)
+Inductive cmd : Set :=
+| cm : trm -> ect -> cmd
+with trm : Set :=
+| tvar : var_trm -> trm
+| lam : var_trm -> trm -> trm
+| mu : var_ect -> cmd -> trm
+| injl : trm -> trm
+| injr : trm -> trm
+| paire : trm -> trm -> trm
+with ect : Set :=
+| evar : var_ect -> ect
+| mut : var_trm -> cmd -> ect
+| app : trm -> ect -> ect
+| disj : ect -> ect -> ect
+| projl : ect -> ect
+| projr : ect -> ect.
+
+Parameter var_formula : Set.
+
+Inductive formula : Set :=
+| Atom : var_formula -> formula
+| Imp : formula -> formula -> formula
+| Disj : formula -> formula -> formula
+| Conj : formula -> formula -> formula.
+
+Let cxt_trm := list (var_trm*formula).
+Let cxt_ect := list (var_ect*formula).
+
+Reserved Notation "c : ( Gamma |- Delta ) " (at level 70).
+Reserved Notation "Gamma |- ( t : A ) ; Delta"
+ (at level 70, A at next level, t at next level).
+Reserved Notation "Gamma ; ( e : A ) |- Delta"
+ (at level 70, A at next level, e at next level).
+
+(** The LK-mu-mu-tilde sequent calculus *)
+Inductive proof_cmd : cmd -> cxt_trm -> cxt_ect -> Set :=
+| Cut v e Gamma Delta A :
+ proof_trm Gamma v A Delta ->
+ proof_ect Gamma e A Delta ->
+ proof_cmd (cm v e) Gamma Delta
+
+where "c : ( Gamma |- Delta )" := (proof_cmd c Gamma Delta)
+
+with proof_trm : cxt_trm -> trm -> formula -> cxt_ect -> Set :=
+| AxR Gamma x A Delta :
+ In (x,A) Gamma ->
+ proof_trm Gamma (tvar x) A Delta
+
+| Mu Gamma alpha c A Delta :
+ proof_cmd c Gamma ((alpha,A)::Delta) ->
+ proof_trm Gamma (mu alpha c) A Delta
+
+| ImpR Gamma x t A B Delta :
+ proof_trm ((x,A)::Gamma) t B Delta ->
+ proof_trm Gamma (lam x t) (Imp A B) Delta
+
+| DisjRL Gamma v A1 A2 Delta :
+ proof_trm Gamma v A1 Delta ->
+ proof_trm Gamma (injl v) (Disj A1 A2) Delta
+
+| DisjRR Gamma v A1 A2 Delta :
+ proof_trm Gamma v A2 Delta ->
+ proof_trm Gamma (injr v) (Disj A1 A2) Delta
+
+| ConjR Gamma v1 v2 A1 A2 Delta :
+ proof_trm Gamma v1 A1 Delta ->
+ proof_trm Gamma v2 A2 Delta ->
+ proof_trm Gamma (paire v1 v2) (Conj A1 A2) Delta
+
+where "Gamma |- ( t : A ) ; Delta" := (proof_trm Gamma t A Delta)
+
+with proof_ect : cxt_trm -> ect -> formula -> cxt_ect -> Set :=
+| AxL Gamma alpha A Delta :
+ In (alpha,A) Delta ->
+ proof_ect Gamma (evar alpha) A Delta
+
+| MuT Gamma x c A Delta :
+ proof_cmd c ((x,A)::Gamma) Delta ->
+ proof_ect Gamma (mut x c) A Delta
+
+| ImpL Gamma v e A B Delta :
+ proof_trm Gamma v A Delta ->
+ proof_ect Gamma e B Delta ->
+ proof_ect Gamma (app v e) (Imp A B) Delta
+
+| DisjL Gamma e1 e2 A1 A2 Delta :
+ proof_ect Gamma e1 A1 Delta ->
+ proof_ect Gamma e2 A2 Delta ->
+ proof_ect Gamma (disj e1 e2) (Disj A1 A2) Delta
+
+| ConjLL Gamma e A1 A2 Delta :
+ proof_ect Gamma e A1 Delta ->
+ proof_ect Gamma (projl e) (Conj A1 A2) Delta
+
+| ConjLR Gamma e A1 A2 Delta :
+ proof_ect Gamma e A2 Delta ->
+ proof_ect Gamma (projr e) (Conj A1 A2) Delta
+
+where "Gamma ; ( e : A ) |- Delta" := (proof_ect Gamma e A Delta)
+.
+
+Section Abstract_Semantics.
+ (** An abstract Kripke-style structure: [wle] is the preorder,
+ [exploding] is the exploding-world predicate, [aforces] is strong
+ refutation of atomic formulae. *)
+ Record Kripke : Type := {
+ world :> Set;
+ wle : world -> world -> Type;
+ wle_refl : forall w, wle w w;
+ wle_trans : forall w w' w'', wle w w' -> wle w' w'' -> wle w w'';
+ exploding : world -> Set;
+ arefutes : world -> var_formula -> Set;
+ arefutes_mon : forall w X, arefutes w X ->
+ forall w', wle w w' -> arefutes w' X
+ }.
+
+ Notation "w <= w'" := (wle w w').
+
+ Variable K:Kripke.
+
+ Fixpoint srefutes (w:K)(A:formula) {struct A} : Type :=
+ match A with
+ | Atom X => arefutes w X
+ | Imp A1 A2 => prod
+ (forall w', w <= w' -> srefutes w' A1 -> exploding w')
+ (forall w', w <= w' ->
+ (forall w'', w' <= w'' -> srefutes w'' A2 -> exploding w'')
+ -> exploding w')
+ | Disj A1 A2 => prod
+ (forall w', w <= w' ->
+ (forall w'', w' <= w'' -> srefutes w'' A1 -> exploding w'')
+ -> exploding w')
+ (forall w', w <= w' ->
+ (forall w'', w' <= w'' -> srefutes w'' A2 -> exploding w'')
+ -> exploding w')
+ | Conj A1 A2 => sum
+ (forall w', w <= w' ->
+ (forall w'', w' <= w'' -> srefutes w'' A1 -> exploding w'')
+ -> exploding w')
+ (forall w', w <= w' ->
+ (forall w'', w' <= w'' -> srefutes w'' A2 -> exploding w'')
+ -> exploding w')
+ end.
+
+ Notation "w : | A ||- " := (srefutes w A) (at level 70).
+
+ Definition forces (w:K)(A:formula) :=
+ forall w':K, w <= w' -> srefutes w' A -> exploding w'.
+
+ Notation "w : ||- A" := (forces w A) (at level 70).
+
+ Definition refutes (w:K)(A:formula) :=
+ forall w', w <= w' -> forces w' A -> exploding w'.
+
+ Notation "w : A ||- " := (refutes w A) (at level 70).
+
+ Fixpoint forces_cxt (w:K)(Gamma:cxt_trm) : Type :=
+ match Gamma with
+ | nil => unit
+ | cons xA Gamma' => prod (w:||-(snd xA)) (forces_cxt w Gamma')
+ end.
+
+ Fixpoint refutes_cxt (w:K)(Delta:cxt_ect) : Type :=
+ match Delta with
+ | nil => unit
+ | cons aA Delta' => prod (refutes w (snd aA)) (refutes_cxt w Delta')
+ end.
+
+ Fixpoint srefutes_cxt (w:K)(Delta:cxt_ect) : Type :=
+ match Delta with
+ | nil => unit
+ | cons aA Delta' => prod (srefutes w (snd aA)) (srefutes_cxt w Delta')
+ end.
+
+ Notation " w : ||-- Gamma" := (forces_cxt w Gamma) (at level 70).
+ Notation " w : Delta ||-- " := (refutes_cxt w Delta) (at level 70).
+
+ (* Combined Scheme only works for sort Prop, so I have to give the
+ definition by hand by adapting a copy-paste of the Prop version: *)
+(* Scheme proof_cmd_rect' := Induction for proof_cmd Sort Prop *)
+(* with proof_trm_rect' := Induction for proof_trm Sort Prop *)
+(* with proof_ect_rect' := Induction for proof_ect Sort Prop. *)
+(* Combined Scheme proof_rect'' from proof_cmd_rect', proof_trm_rect', *)
+(* proof_ect_rect'. *)
+(* Unset Printing Notations. *)
+(* Print proof_rect''. *)
+(* Set Printing Notations. *)
+
+ Scheme proof_cmd_rect' := Induction for proof_cmd Sort Type
+ with proof_trm_rect' := Induction for proof_trm Sort Type
+ with proof_ect_rect' := Induction for proof_ect Sort Type.
+
+ (** We have to make the mutual-induction predicate "by hand" since the Combined Scheme Coq command works only in sort Prop, not in Type. *)
+ Definition proof_rect' :=
+ fun
+ (P : forall (c : cmd) (c0 : cxt_trm) (c1 : cxt_ect),
+ proof_cmd c c0 c1 -> Type)
+ (P0 : forall (c : cxt_trm) (t : trm) (t0 : formula) (c0 : cxt_ect),
+ proof_trm c t t0 c0 -> Type)
+ (P1 : forall (c : cxt_trm) (e : ect) (t : formula) (c0 : cxt_ect),
+ proof_ect c e t c0 -> Type)
+ (f : forall (v : trm) (e : ect) (Gamma : cxt_trm)
+ (Delta : cxt_ect) (A : formula) (p : proof_trm Gamma v A Delta),
+ P0 Gamma v A Delta p ->
+ forall p0 : proof_ect Gamma e A Delta,
+ P1 Gamma e A Delta p0 -> P (cm v e) Gamma Delta (Cut p p0))
+ (f0 : forall (Gamma : list (prod var_trm formula)) (x : var_trm)
+ (A : formula) (Delta : cxt_ect) (i : In (pair x A) Gamma),
+ P0 Gamma (tvar x) A Delta (AxR (Gamma:=Gamma) (x:=x) (A:=A) Delta i))
+ (f1 : forall (Gamma : cxt_trm) (alpha : var_ect)
+ (c : cmd) (A : formula) (Delta : list (prod var_ect formula))
+ (p : proof_cmd c Gamma (cons (pair alpha A) Delta)),
+ P c Gamma (cons (pair alpha A) Delta) p ->
+ P0 Gamma (mu alpha c) A Delta (Mu p))
+ (f2 : forall (Gamma : list (prod var_trm formula)) (x : var_trm)
+ (t : trm) (A B : formula) (Delta : cxt_ect)
+ (p : proof_trm (cons (pair x A) Gamma) t B Delta),
+ P0 (cons (pair x A) Gamma) t B Delta p ->
+ P0 Gamma (lam x t) (Imp A B) Delta (ImpR p))
+ (f3 : forall (Gamma : cxt_trm) (v : trm) (A1 A2 : formula)
+ (Delta : cxt_ect) (p : proof_trm Gamma v A1 Delta),
+ P0 Gamma v A1 Delta p ->
+ P0 Gamma (injl v) (Disj A1 A2) Delta (DisjRL A2 p))
+ (f4 : forall (Gamma : cxt_trm) (v : trm) (A1 A2 : formula)
+ (Delta : cxt_ect) (p : proof_trm Gamma v A2 Delta),
+ P0 Gamma v A2 Delta p ->
+ P0 Gamma (injr v) (Disj A1 A2) Delta (DisjRR A1 p))
+ (f5 : forall (Gamma : cxt_trm) (v1 v2 : trm) (A1 A2 : formula)
+ (Delta : cxt_ect) (p : proof_trm Gamma v1 A1 Delta),
+ P0 Gamma v1 A1 Delta p ->
+ forall p0 : proof_trm Gamma v2 A2 Delta,
+ P0 Gamma v2 A2 Delta p0 ->
+ P0 Gamma (paire v1 v2) (Conj A1 A2) Delta (ConjR p p0))
+ (f6 : forall (Gamma : cxt_trm) (alpha : var_ect)
+ (A : formula) (Delta : list (prod var_ect formula))
+ (i : In (pair alpha A) Delta),
+ P1 Gamma (evar alpha) A Delta
+ (AxL Gamma (alpha:=alpha) (A:=A) (Delta:=Delta) i))
+ (f7 : forall (Gamma : list (prod var_trm formula)) (x : var_trm)
+ (c : cmd) (A : formula) (Delta : cxt_ect)
+ (p : proof_cmd c (cons (pair x A) Gamma) Delta),
+ P c (cons (pair x A) Gamma) Delta p ->
+ P1 Gamma (mut x c) A Delta (MuT p))
+ (f8 : forall (Gamma : cxt_trm) (v : trm) (e : ect)
+ (A B : formula) (Delta : cxt_ect) (p : proof_trm Gamma v A Delta),
+ P0 Gamma v A Delta p ->
+ forall p0 : proof_ect Gamma e B Delta,
+ P1 Gamma e B Delta p0 ->
+ P1 Gamma (app v e) (Imp A B) Delta (ImpL p p0))
+ (f9 : forall (Gamma : cxt_trm) (e1 e2 : ect) (A1 A2 : formula)
+ (Delta : cxt_ect) (p : proof_ect Gamma e1 A1 Delta),
+ P1 Gamma e1 A1 Delta p ->
+ forall p0 : proof_ect Gamma e2 A2 Delta,
+ P1 Gamma e2 A2 Delta p0 ->
+ P1 Gamma (disj e1 e2) (Disj A1 A2) Delta (DisjL p p0))
+ (f10 : forall (Gamma : cxt_trm) (e : ect) (A1 A2 : formula)
+ (Delta : cxt_ect) (p : proof_ect Gamma e A1 Delta),
+ P1 Gamma e A1 Delta p ->
+ P1 Gamma (projl e) (Conj A1 A2) Delta (ConjLL A2 p))
+ (f11 : forall (Gamma : cxt_trm) (e : ect) (A1 A2 : formula)
+ (Delta : cxt_ect) (p : proof_ect Gamma e A2 Delta),
+ P1 Gamma e A2 Delta p ->
+ P1 Gamma (projr e) (Conj A1 A2) Delta (ConjLR A1 p)) =>
+ pair (pair
+ (proof_cmd_rect' (P:=P) (P0:=P0) (P1:=P1) f f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11)
+ (proof_trm_rect' (P:=P) (P0:=P0) (P1:=P1) f f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11))
+ (proof_ect_rect' (P:=P) (P0:=P0) (P1:=P1) f f0 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11).
+
+ Lemma forces_cxt_In : forall w x A Gamma,
+ In (x, A) Gamma -> w:||--Gamma -> w:||-A.
+ Proof.
+ induction Gamma.
+ simpl.
+ intros.
+ contradict H.
+ simpl.
+ intros H H0.
+ destruct H.
+ rewrite e in H0.
+ intuition.
+ intuition.
+ Defined.
+
+ Lemma refutes_cxt_In : forall w alpha A Delta,
+ In (alpha, A) Delta -> w:Delta||-- -> w:A||-.
+ Proof.
+ induction Delta.
+ simpl.
+ intros.
+ contradict H.
+ simpl.
+ intros H H0.
+ destruct H.
+ rewrite e in H0.
+ intuition.
+ intuition.
+ Defined.
+
+ Lemma srefutes_mon : forall A w, w:|A||- -> forall w', w <= w' -> w':|A||-.
+ Proof.
+ induction A.
+ simpl.
+ intros.
+ eauto using arefutes_mon.
+
+ intros w H1 w' ww'.
+ simpl.
+ simpl in H1.
+ destruct H1 as [H11 H12].
+ split.
+ eauto using wle_trans.
+ eauto using wle_trans.
+
+ intros w H1 w' ww'.
+ simpl in *.
+ destruct H1.
+ split.
+ eauto using wle_trans.
+ eauto using wle_trans.
+
+ intros w H1 w' ww'.
+ simpl in *.
+ destruct H1.
+ eauto using wle_trans.
+ eauto using wle_trans.
+ Defined.
+
+ Lemma forces_Imp_r : forall (w:K)(A B:formula),
+ (forall w', w <= w' -> w':||-A -> w':||-B) -> w:||- (Imp A B).
+ Proof.
+ intros w A B H.
+ unfold forces in *.
+ simpl in *.
+ intros w' ww' [H1 H2].
+ apply H2.
+ apply wle_refl.
+ intros w'' w'w'' H3.
+ apply H with w'.
+ assumption.
+ assumption.
+ assumption.
+ assumption.
+ Defined.
+
+ Lemma forces_Imp_l : forall (w:K)(A B:formula),
+ w:||- (Imp A B) -> forall w', w <= w' -> w':||-A -> w':||-B.
+ Proof.
+ intros w A B H1 w' ww' H2.
+ unfold forces in H1.
+ unfold forces.
+ intros w'' w'w'' H3.
+ apply H1.
+ eauto using wle_trans.
+ simpl.
+ split.
+ intros w''' w''w''' H4.
+ apply H2; eauto using wle_trans.
+ intros w''' w''w''' H4.
+ apply H4.
+ apply wle_refl.
+ eauto using wle_trans,srefutes_mon.
+ Defined.
+
+ Lemma forces_mon : forall A w, w:||-A -> forall w', w <= w' -> w':||-A.
+ Proof.
+ induction A.
+ intros.
+ unfold forces in *.
+ eauto using wle_trans.
+
+ intros w H1 w' ww'.
+ unfold forces in *.
+ eauto using wle_trans.
+
+ unfold forces.
+ simpl.
+ intros w1 H1 w2 H2.
+ intros w3 H3 H4.
+ apply H1.
+ eauto using wle_trans.
+ destruct H4.
+ intuition.
+
+ unfold forces; simpl.
+ intros w1 H1 w2 H2.
+ intros w3 H3 H4.
+ apply H1.
+ eauto using wle_trans.
+ destruct H4.
+ intuition.
+ intuition.
+ Defined.
+
+ Lemma forces_cxt_mon :
+ forall Gamma w, w:||--Gamma -> forall w', w <= w' -> w':||--Gamma.
+ Proof.
+ induction Gamma.
+ simpl.
+ auto.
+ simpl.
+ intros.
+ destruct X.
+ split; eauto using wle_trans,forces_mon.
+ Defined.
+
+ Lemma srefutes_refutes (w:K)(A:formula) : w:|A||- -> w:A||-.
+ Proof.
+ intros H.
+ unfold refutes.
+ intros w' ww' H0.
+ unfold forces in H0.
+ eauto using wle_refl, srefutes_mon.
+ Defined.
+
+ Lemma refutes_mon : forall A w, w:A||- -> forall w', w <= w' -> w':A||-.
+ Proof.
+ induction A.
+ intros.
+ unfold refutes in *.
+ unfold forces in *.
+ eauto using wle_trans.
+
+ intros w H1 w' ww'.
+ unfold refutes in *.
+ unfold forces in *.
+ eauto using wle_trans.
+
+ intros w H1 w' ww'.
+ unfold refutes in *.
+ unfold forces in *.
+ eauto using wle_trans.
+
+ intros w H1 w' ww'.
+ unfold refutes in *.
+ unfold forces in *.
+ eauto using wle_trans.
+ Defined.
+
+ Lemma refutes_cxt_mon :
+ forall Delta w, w:Delta||-- -> forall w', w <= w' -> w':Delta||--.
+ Proof.
+ induction Delta.
+ simpl.
+ auto.
+ simpl.
+ intros.
+ destruct X.
+ split; eauto using wle_trans,refutes_mon.
+ Defined.
+
+ Theorem soundness :
+ (forall c Gamma Delta, c:(Gamma|-Delta) ->
+ forall w, w:||--Gamma -> w:Delta||-- -> exploding w) *
+
+ (forall Gamma t A Delta, Gamma|-(t:A);Delta ->
+ forall w, w:||--Gamma -> w:Delta||-- -> w:||-A) *
+
+ (forall Gamma e A Delta, Gamma;(e:A)|-Delta ->
+ forall w, w:||--Gamma -> w:Delta||-- -> w:A||-).
+ Proof.
+ apply proof_rect'.
+
+ (* Cut *)
+ intros v e Gamma Delta A _ IH1 _ IH2.
+ intros w H H0.
+ apply IH2 with w; auto.
+ apply wle_refl.
+
+ (* Ax_R *)
+ intros Gamma x A Delta IH1.
+ intros w H H0 w' ww' H1.
+ assert (H2:=forces_cxt_In IH1 H).
+ auto.
+
+ (* mu *)
+ intros Gamma alpha c A Delta _ IH1 w H H0.
+ intros w' ww' H1.
+ apply IH1.
+ eauto using forces_cxt_mon,wle_trans.
+ simpl.
+ split; eauto using refutes_cxt_mon,refutes_mon,wle_trans,srefutes_refutes.
+
+ (* imp_R *)
+ intros Gamma x t A B Delta _ H.
+ simpl in H.
+ intros w H0 H1.
+ apply forces_Imp_r.
+ intros w' ww' H2.
+ apply H.
+ split; eauto using forces_cxt_mon,forces_mon,wle_trans.
+ eauto using refutes_cxt_mon,wle_trans.
+
+ (* DisjRL *)
+ intros Gamma v A1 A2 Delta _ IH1.
+ intros w H H0.
+ red.
+ intros w' H1 H2.
+ simpl in H2.
+ destruct H2 as [H21 H22].
+ apply H21.
+ apply wle_refl.
+ apply IH1;
+ eauto using forces_cxt_mon,refutes_cxt_mon.
+
+ (* DisjRR *)
+ intros Gamma v A1 A2 Delta _ IH1.
+ intros w H H0.
+ red.
+ intros w' H1 H2.
+ simpl in H2.
+ destruct H2 as [H21 H22].
+ apply H22.
+ apply wle_refl.
+ apply IH1;
+ eauto using forces_cxt_mon,refutes_cxt_mon.
+
+ (* ConjR *)
+ intros Gamma v1 v2 A1 A2 Delta _ IH1 _ IH2.
+ intros w H H0.
+ red.
+ intros w' H1 H2.
+ simpl in H2.
+ destruct H2.
+ apply e.
+ apply wle_refl.
+ apply IH1;
+ eauto using forces_cxt_mon,refutes_cxt_mon.
+ apply e.
+ apply wle_refl.
+ apply IH2;
+ eauto using forces_cxt_mon,refutes_cxt_mon.
+
+ (* Ax_L *)
+ intros Gamma alpha A Delta IH1.
+ intros w H H0.
+ assert (H1 := refutes_cxt_In IH1 H0).
+ assumption.
+
+ (* mu-tilda *)
+ intros Gamma x c A Delta _ IH1.
+ intros w H H0.
+ simpl in IH1.
+ unfold refutes.
+ intros w' ww' H1.
+ apply IH1.
+ split; eauto using wle_trans,forces_mon,forces_cxt_mon.
+ eauto using wle_trans,refutes_cxt_mon.
+
+ (* imp_L *)
+ intros Gamma v e A B Delta _ IH1 _ IH2 w HGamma HDelta.
+ unfold refutes.
+ intros w' ww' H.
+ apply H.
+ apply wle_refl.
+ change (forces w' A * refutes w' B).
+ split;eauto using wle_trans,forces_cxt_mon,refutes_cxt_mon.
+
+ (* DisjL *)
+ intros Gamma e1 e2 A1 A2 Delta _ IH1 _ IH2.
+ intros w H H0.
+ red.
+ intros w' H1 H2.
+ apply H2.
+ apply wle_refl.
+ simpl.
+ split.
+ apply IH1.
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+ apply IH2.
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+
+ (* ConjLL *)
+ intros Gamma e A1 A2 Delta _ IH w H H0.
+ intros w' H1 H2.
+ apply H2.
+ apply wle_refl.
+ simpl.
+ left.
+ apply IH;
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+
+ (* ConjLR *)
+ intros Gamma e A1 A2 Delta _ IH w H H0.
+ intros w' H1 H2.
+ apply H2.
+ apply wle_refl.
+ simpl.
+ right.
+ apply IH;
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+ Defined.
+End Abstract_Semantics.
+
+(** The Universal model *)
+Section Context_Semantics.
+ Definition Kworld : Set := cxt_trm*cxt_ect.
+
+ Definition Kle (w w':Kworld) : Type :=
+ incl (fst w) (fst w') * incl (snd w) (snd w').
+
+ Lemma Kle_refl : forall w, Kle w w.
+ Proof.
+ intro w.
+ split;
+ auto using incl_refl.
+ Defined.
+
+ Lemma Kle_trans : forall w w' w'', Kle w w' -> Kle w' w'' -> Kle w w''.
+ Proof.
+ unfold Kle.
+ intros.
+ intuition eauto using incl_tran.
+ Defined.
+
+ Definition Kexploding (w:Kworld) : Set :=
+ {c:cmd & proof_cmd c (fst w) (snd w)}.
+
+ Definition Karefutes (w:Kworld)(X:var_formula) : Set :=
+ {e:ect & proof_ect (fst w) e (Atom X) (snd w)}.
+
+ Lemma proof_weaken :
+ (forall c Gamma Delta, c:(Gamma|-Delta) ->
+ forall Gamma' Delta', incl Gamma Gamma' -> incl Delta Delta' ->
+ c:(Gamma'|-Delta')) *
+ (forall Gamma t A Delta, Gamma|-(t:A);Delta ->
+ forall Gamma' Delta', incl Gamma Gamma' -> incl Delta Delta' ->
+ Gamma'|-(t:A);Delta') *
+ (forall Gamma e A Delta, Gamma;(e:A)|-Delta ->
+ forall Gamma' Delta', incl Gamma Gamma' -> incl Delta Delta' ->
+ Gamma';(e:A)|-Delta').
+ Proof.
+ apply proof_rect'.
+
+ (* Cut *)
+ eauto using Cut.
+
+ (* AxR *)
+ eauto using AxR.
+
+ (* Mu *)
+ intros.
+ apply Mu.
+ apply X.
+ assumption.
+ red.
+ intuition.
+ simpl in H |- *.
+ intuition.
+
+ (* ImpR *)
+ intros.
+ apply ImpR.
+ apply X.
+ red.
+ simpl.
+ intuition.
+ assumption.
+
+ (* DisjRL *)
+ intros.
+ apply DisjRL.
+ apply X.
+ assumption.
+ assumption.
+
+ (* DisjRR *)
+ intros.
+ apply DisjRR.
+ apply X.
+ assumption.
+ assumption.
+
+ (* ConjR *)
+ intros.
+ apply ConjR.
+ apply X.
+ assumption.
+ assumption.
+ apply X0.
+ assumption.
+ assumption.
+
+ (* AxL *)
+ eauto using AxL.
+
+ (* MuT *)
+ intros.
+ apply MuT.
+ apply X.
+ red;simpl;intuition.
+ assumption.
+
+ (* ImpL *)
+ eauto using ImpL.
+
+ (* DisjL *)
+ intros.
+ apply DisjL.
+ apply X; intuition.
+ apply X0; intuition.
+
+ (* ConjLL *)
+ intros.
+ apply ConjLL.
+ apply X; intuition.
+
+ (* ConjLR *)
+ intros.
+ apply ConjLR.
+ apply X; intuition.
+ Defined.
+
+ Lemma Karefutes_mon : forall w X, Karefutes w X ->
+ forall w', Kle w w' -> Karefutes w' X.
+ Proof.
+ intros w X H.
+ destruct H as [e He].
+ inversion He.
+
+ intros w' Hle.
+ exists (evar alpha).
+ apply AxL.
+ clear -H Hle.
+ destruct Hle.
+ intuition.
+
+ intros w' Hle.
+ exists (mut x c).
+ apply MuT.
+ destruct Hle.
+ eapply (fst (fst proof_weaken)).
+ apply H.
+ red;simpl;intuition.
+ assumption.
+ Defined.
+
+ Definition K : Kripke.
+ (* begin show *)
+ apply Build_Kripke with Kworld Kle Karefutes.
+ exact Kle_refl.
+ exact Kle_trans.
+ exact Kexploding.
+ exact Karefutes_mon.
+ (* end show *)
+ Defined.
+
+ Definition Neutral_trm (uA:var_trm*formula)(Gamma:cxt_trm)(Delta:cxt_ect) :=
+ forall t C Gamma' Delta', incl Gamma Gamma' -> incl Delta Delta' ->
+ (uA::Gamma')|-(t:C);Delta' -> Gamma'|-(t:C);Delta'.
+ Hint Unfold Neutral_trm.
+
+ Definition Neutral_ect (alphaA:var_ect*formula)(Gamma:cxt_trm)(Delta:cxt_ect) :=
+ forall e C Gamma' Delta', incl Gamma Gamma' -> incl Delta Delta' ->
+ Gamma';(e:C)|-(alphaA::Delta') -> Gamma';(e:C)|-Delta'.
+ Hint Unfold Neutral_ect.
+
+ Notation "w <= w'" := (Kle w w').
+
+ Notation "w : | A ||- " := (@srefutes K w A) (at level 70).
+ Notation "w : ||- A" := (@forces K w A) (at level 70).
+ Notation "w : A ||- " := (@refutes K w A) (at level 70).
+ Notation " w : ||-- Gamma" := (@forces_cxt K w Gamma) (at level 70).
+ Notation " w : Delta ||-- " := (@refutes_cxt K w Delta) (at level 70).
+
+ (** We use these variables as a supply of fresh ones. This was just a temporary solution, for total correctness I should include an explicit counter to Kcompleteness, like it is done in the other files for Kripke-style models. *)
+ Parameter beta gamma delta:var_ect.
+ Parameter y:var_trm.
+
+ (** The proof of completeness is via a two pairs of reify-reflect function *)
+ Theorem Kcompleteness : forall A Gamma Delta,
+ ((Gamma,Delta):||-A -> sigT (fun t:trm => Gamma|-(t:A);Delta)) *
+ (forall x, Neutral_trm (x,A) Gamma Delta -> (Gamma,Delta):||-A) *
+ ((Gamma,Delta):A||- -> sigT (fun e:ect => Gamma;(e:A)|-Delta)) *
+ (forall alpha, Neutral_ect (alpha,A) Gamma Delta -> (Gamma,Delta):A||-).
+ Proof.
+ induction A.
+
+ (* Atom type *)
+ intuition.
+
+ unfold forces in X.
+ simpl in X.
+ unfold Kexploding in X.
+ unfold Karefutes in X.
+ set (Delta' := ((beta, Atom v)::Delta)).
+ assert (H : (Gamma,Delta) <= (Gamma,Delta')).
+ unfold Kle.
+ split.
+ intuition.
+ unfold Delta'.
+ simpl.
+ red.
+ simpl.
+ intuition.
+ assert (X0 := X (Gamma,Delta') H).
+ simpl in X0.
+ clear -X0.
+ assert (sigT (fun (e : ect) => Gamma; (e : Atom v)|-Delta')).
+ exists (evar beta).
+ apply AxL.
+ unfold Delta'.
+ simpl.
+ intuition.
+ destruct (X0 H) as [c Hc].
+ exists (mu beta c).
+ apply Mu.
+ apply Hc.
+
+ rename X into H.
+ red in H.
+ red.
+ simpl.
+ red.
+ unfold Karefutes.
+ intros w' H0 H1.
+ assert (sigT (fun (Gamma':cxt_trm) =>
+ sigT (fun (Delta':cxt_ect) => w'=(Gamma',Delta')))).
+ exists (fst w').
+ exists (snd w').
+ induction w'.
+ simpl.
+ reflexivity.
+ destruct H2 as [Gamma' [Delta' Hw']].
+ rewrite Hw' in *.
+ simpl in *.
+ destruct H1 as [e He].
+ assert (H1 := H (tvar x) (Atom v) Gamma' Delta).
+ exists (cm (tvar x) e).
+ apply Cut with (Atom v).
+ apply (snd (fst proof_weaken)) with Gamma' Delta; destruct H0; intuition.
+ apply X.
+ apply incl_refl.
+ apply (AxR).
+ simpl;intuition.
+ assumption.
+
+ unfold refutes in X.
+ simpl in X.
+ unfold forces in X.
+ simpl in X.
+ unfold Kexploding in X.
+ set (w' := ((y,(Atom v))::Gamma, Delta)).
+ assert ((Gamma,Delta) <= w').
+ unfold w'.
+ red.
+ simpl.
+ intuition.
+ assert (forall w'' : Kworld, w' <= w'' -> Karefutes w'' v ->
+ sigT (fun (c:cmd) => c : (fst w'' |- snd w''))).
+ clear.
+ unfold Karefutes.
+ intros w'' H H0.
+ destruct H0 as [e He].
+ exists (cm (tvar y) e).
+ apply Cut with (Atom v).
+ apply AxR.
+ destruct H.
+ unfold w' in i.
+ clear -i.
+ red in i.
+ red.
+ apply i.
+ simpl.
+ intuition.
+ assumption.
+ destruct (X w' X0 X1) as [c Hc].
+ exists (mut y c).
+ apply MuT.
+ apply Hc.
+
+ red in X.
+ red.
+ simpl.
+ red.
+ intro w'.
+ unfold forces.
+ simpl.
+ unfold Karefutes, Kexploding .
+ assert (sigT (fun (Gamma':cxt_trm) =>
+ sigT (fun (Delta':cxt_ect) => w'=(Gamma',Delta')))).
+ exists (fst w').
+ exists (snd w').
+ induction w'.
+ simpl.
+ reflexivity.
+ destruct H as [Gamma' [Delta' Hw']].
+ rewrite Hw' in *.
+ simpl in *.
+ intros.
+ replace Gamma' with (fst w').
+ replace Delta' with (snd w').
+ Focus 2.
+ rewrite Hw'; auto.
+ Focus 2.
+ rewrite Hw'; auto.
+ apply X1.
+ red.
+ rewrite Hw'.
+ intuition.
+ exists (evar alpha).
+ apply (snd proof_weaken) with Gamma Delta'.
+ Focus 2.
+ rewrite Hw'.
+ simpl.
+ destruct X0.
+ intuition.
+ Focus 2.
+ rewrite Hw'.
+ simpl.
+ intuition.
+ apply X.
+ destruct X0.
+ intuition.
+ destruct X0.
+ intuition.
+ apply AxL.
+ simpl.
+ intuition.
+
+ (* Induction case -- Imp *)
+ intuition.
+ unfold forces in X.
+ simpl exploding in X.
+ simpl wle in X.
+ unfold Kexploding in X.
+ set (w':=((y,A1)::Gamma,(beta,A2)::Delta)).
+ assert ((Gamma,Delta) <= w').
+ unfold w'.
+ red.
+ simpl.
+ split;
+ red;
+ simpl;
+ intuition.
+ assert(H0 : w' : | Imp A1 A2 ||-).
+ simpl.
+ split.
+ (* prove: forces w' A1 *)
+ clear -IHA1.
+ destruct (IHA1 (fst w') (snd w')).
+ destruct p.
+ destruct p.
+ unfold forces in f.
+ simpl in f.
+ unfold w'.
+ apply f with y.
+ clear.
+ red.
+ intros.
+ apply (snd (fst proof_weaken)) with ((y, A1) :: Gamma') Delta'.
+ assumption.
+ red in X |- *;simpl in *;intuition.
+ red in X0 |- *;simpl in *;intuition.
+ (* prove: refutes w' A2 *)
+ clear -IHA2.
+ destruct (IHA2 (fst w') (snd w')).
+ destruct p.
+ destruct p.
+ unfold refutes in r.
+ simpl in r.
+ unfold forces in r.
+ simpl in r.
+ unfold w'.
+ apply r with beta.
+ clear.
+ red.
+ intros.
+ apply (snd proof_weaken) with Gamma' ((beta,A2)::Delta').
+ assumption.
+ intuition.
+ clear -X0; red in X0 |- *; simpl in *; intuition.
+ destruct (X w' X0 H0) as [c Hc].
+ unfold w' in Hc.
+ simpl in Hc.
+ eexists.
+ apply ImpR.
+ apply Mu.
+ apply Hc.
+
+ unfold forces.
+ simpl.
+ intros w' H H0.
+ unfold Kexploding.
+ destruct H0 as [X1 X2].
+ assert (H1 : w':||-A1).
+ unfold forces.
+ apply X1.
+ assert (H2 : w':A2||-).
+ unfold refutes.
+ apply X2.
+ assert (IH1 := fst (fst (fst (IHA1 (fst w') (snd w')))) H1).
+ assert (IH2 := (snd (fst (IHA2 (fst w') (snd w')))) H2).
+ clear -IH1 IH2 H X.
+ destruct IH1 as [t Ht].
+ destruct IH2 as [e He].
+ assert (Hu : sigT (fun u:trm => fst w' |- (u:A2); snd w')).
+ red in X.
+ exists (mu beta (cm (tvar x) (app t e))).
+ apply X.
+ destruct H; intuition.
+ destruct H; intuition.
+ apply Mu.
+ apply Cut with (Imp A1 A2).
+ apply AxR.
+ red; simpl; intuition.
+ apply ImpL.
+ apply (snd (fst proof_weaken)) with (fst w') (snd w').
+ apply Ht.
+ intuition.
+ intuition.
+ apply (snd (proof_weaken)) with (fst w') (snd w').
+ assumption.
+ intuition.
+ intuition.
+ destruct Hu as [u Hu].
+ exists (cm u e). (* can this be turned into a trivial cut?
+ -- yes, se the paper version, where we
+ don't cut with A2 but with A1->A2 *)
+ apply Cut with A2.
+ assumption.
+ assumption.
+
+ red in X.
+ simpl in X.
+ red in X.
+ set (w1:=(Gamma,(beta,A1)::Delta)).
+ assert ((Gamma,Delta) <= w1).
+ unfold w1.
+ split.
+ simpl.
+ intuition.
+ simpl.
+ intuition.
+ assert (H0 : w1:||-Imp A1 A2).
+ unfold forces.
+ simpl.
+ intros.
+ destruct X0.
+ assert (sigT (fun e:ect => fst w';(e:A1) |- snd w')).
+ exists (evar beta).
+ apply AxL.
+ destruct X1 as [_ H2].
+ apply H2.
+ unfold w1.
+ simpl; intuition.
+ assert (sigT (fun t:trm => fst w' |- (t:A1);snd w')).
+ destruct X2.
+ assert (IH1 := (fst (fst (fst (IHA1 (fst w') (snd w'))))) k).
+ apply IH1.
+ destruct H0 as [t Ht].
+ destruct H as [e He].
+ exists (cm t e).
+ apply Cut with A1.
+ assumption.
+ assumption.
+ set (w2:=((y,A2)::Gamma,Delta)).
+ assert ((Gamma,Delta) <= w2).
+ unfold w2.
+ split.
+ simpl.
+ intuition.
+ simpl.
+ intuition.
+ assert (H2 : w2:||-Imp A1 A2).
+ unfold forces.
+ simpl.
+ intros.
+ destruct X0.
+ assert (sigT (fun t:trm => fst w' |- (t:A2); snd w')).
+ exists (tvar y).
+ apply AxR.
+ destruct X2 as [H2 _].
+ apply H2.
+ unfold w2.
+ simpl; intuition.
+ assert (sigT (fun e:ect => fst w' ; (e:A2) |- snd w')).
+ destruct X3.
+ assert (IH2 := (snd (fst (IHA2 (fst w') (snd w')))) k0).
+ apply IH2.
+ destruct H as [t Ht].
+ destruct H1 as [e He].
+ exists (cm t e).
+ apply Cut with A2.
+ assumption.
+ assumption.
+ destruct (X w1 X0 H0) as [c1 Hc1].
+ destruct (X w2 X1 H2) as [c2 Hc2].
+ unfold w1 in Hc1.
+ simpl in Hc1.
+ unfold w2 in Hc2.
+ simpl in Hc2.
+ eexists.
+ apply ImpL.
+ apply Mu.
+ apply Hc1.
+ apply MuT.
+ apply Hc2.
+
+ unfold refutes.
+ simpl.
+ intros.
+ red.
+ unfold forces in X.
+ simpl exploding in X.
+ simpl wle in X.
+ red in X.
+ set (w'' := ((y,A1)::fst w', (alpha,A2)::snd w')).
+ assert (w' <= w'').
+ unfold w'';simpl;red;simpl;intuition.
+ assert (H2 : w'':|Imp A1 A2||-).
+ simpl.
+ split.
+ (* forces w'' A1 *)
+ assert (IH1 := snd (fst (fst (IHA1 (fst w'') (snd w''))))).
+ unfold forces in IH1.
+ simpl in IH1.
+ apply IH1 with y.
+ clear -X0.
+ red.
+ intros.
+ apply (snd (fst proof_weaken)) with ((y,A1)::Gamma') Delta'.
+ assumption.
+ red in X |- *.
+ simpl in *.
+ intuition.
+ apply incl_refl.
+ (* refutes w'' A2 *)
+ assert (IH2 := snd (IHA2 (fst w'') (snd w''))).
+ unfold refutes in IH2.
+ simpl in IH2.
+ apply IH2 with alpha.
+ clear -X0.
+ red.
+ intros.
+ apply (snd proof_weaken) with Gamma' ((alpha,A2)::Delta').
+ assumption.
+ apply incl_refl.
+ red in X1 |- *.
+ simpl in *.
+ intuition.
+ destruct (X1 w'' X2 H2) as [c Hc].
+ unfold w'' in Hc.
+ simpl in Hc.
+ eexists.
+ apply Cut with (Imp A1 A2).
+ apply ImpR.
+ apply Mu.
+ apply Hc.
+ apply X.
+ red in X0; intuition.
+ red in X0; intuition.
+ apply AxL.
+ red; intuition.
+
+ (* Induction case -- Disj *)
+ split.
+ split.
+ split.
+
+ intro H.
+ unfold forces in H.
+ change (forall w':K, (Gamma,Delta) <= w' -> refutes w' A1 * refutes w' A2 -> Kexploding w') in H.
+ set (w' := (Gamma, (delta,A2)::(gamma,A1)::(beta,Disj A1 A2)::Delta)).
+ assert (H0 : (Gamma,Delta) <= w').
+ unfold w'.
+ clear.
+ red;simpl;split;red;simpl;intuition.
+ assert (H1 : w':A1||-).
+ apply (snd (IHA1 (fst w') (snd w'))) with gamma.
+ unfold w'.
+ simpl.
+ red.
+ intros.
+ apply (snd proof_weaken) with Gamma' ((gamma, A1)::Delta').
+ assumption.
+ apply incl_refl.
+ clear -X0.
+ red in X0 |- *; simpl in X0 |- *;intuition.
+ assert (H2 : w':A2||-).
+ apply (snd (IHA2 (fst w') (snd w'))) with delta.
+ unfold w'.
+ simpl.
+ red.
+ intros.
+ apply (snd proof_weaken) with Gamma' ((delta, A2)::Delta').
+ assumption.
+ apply incl_refl.
+ clear -X0.
+ red in X0 |- *; simpl in X0 |- *;intuition.
+ assert (Hc := H w' H0 (H1,H2)).
+ destruct Hc as [c Hc].
+ unfold w' in Hc.
+ simpl in Hc.
+ exists (
+ mu beta (
+ cm (mu gamma (cm (mu delta c) (
+ mut y (cm (injr (tvar y)) (evar beta)))))
+ (mut y (cm (injl (tvar y)) (evar beta))))).
+ (*eexists.*)
+ apply Mu.
+ apply Cut with A1.
+ apply Mu.
+ apply Cut with A2.
+ apply Mu.
+ apply Hc.
+ apply MuT.
+ apply Cut with (Disj A1 A2).
+ apply DisjRR.
+ apply AxR.
+ simpl;intuition.
+ apply AxL.
+ simpl;intuition.
+ apply MuT.
+ apply Cut with (Disj A1 A2).
+ apply DisjRL.
+ apply AxR.
+ simpl;intuition.
+ apply AxL.
+ simpl;intuition.
+
+ intros x H.
+ unfold forces.
+ change (forall w':K, (Gamma,Delta) <= w' -> refutes w' A1 * refutes w' A2 -> Kexploding w').
+ intros w' H0 [H1 H2].
+ change ((Gamma,Delta) <= (fst w', snd w')) in H0.
+ destruct (IHA1 (fst w') (snd w')) as [[_ IH1] _].
+ apply IH1 in H1.
+ destruct (IHA2 (fst w') (snd w')) as [[_ IH2] _].
+ apply IH2 in H2.
+ destruct H1 as [e1 He1].
+ destruct H2 as [e2 He2].
+ eexists.
+ apply Cut with (Disj A1 A2).
+ apply H.
+ destruct H0.
+ intuition.
+ destruct H0.
+ intuition.
+ apply AxR.
+ intuition.
+ apply DisjL.
+ apply He1.
+ apply He2.
+
+ intro H.
+ unfold refutes in H.
+ unfold forces in H.
+ simpl in H.
+ change (forall w':K, (Gamma,Delta) <= w' ->
+ (forall w'':K, w' <= w'' -> refutes w'' A1 * refutes w'' A2 -> Kexploding w'')
+ -> Kexploding w') in H.
+ set (w' := (((y,Disj A1 A2)::Gamma), Delta)).
+ assert (H0 : (Gamma, Delta) <= w').
+ unfold w'.
+ red.
+ simpl.
+ intuition.
+ apply H in H0.
+ destruct H0 as [c Hc].
+ eexists.
+ apply MuT.
+ unfold w' in Hc.
+ simpl in Hc.
+ apply Hc.
+ intros w'' Hle H1.
+ destruct H1 as [H1 H2].
+ destruct (IHA1 (fst w'') (snd w'')) as [[_ IH1] _].
+ apply IH1 in H1.
+ destruct (IHA2 (fst w'') (snd w'')) as [[_ IH2] _].
+ apply IH2 in H2.
+ destruct H1 as [e1 He1].
+ destruct H2 as [e2 He2].
+ eexists.
+ apply Cut with (Disj A1 A2).
+ apply AxR.
+ destruct Hle as [Hle _].
+ apply Hle.
+ unfold w'.
+ simpl.
+ left.
+ reflexivity.
+ apply DisjL.
+ apply He1.
+ apply He2.
+
+ intros alpha H.
+ unfold refutes.
+ unfold forces.
+ simpl.
+ change (forall w':K, (Gamma,Delta) <= w' ->
+ (forall w'':K, w' <= w'' -> refutes w'' A1 * refutes w'' A2 -> Kexploding w'')
+ -> Kexploding w').
+ intros w' H0 H1.
+ set (w1 := (fst w', ((gamma,A2)::(beta,A1)::(snd w')))).
+ assert (H2 : w' <= w1).
+ change ((fst w',snd w') <= w1).
+ unfold w1.
+ red; simpl;
+ intuition.
+ assert (H3: w1:A1||-).
+ apply (snd (IHA1 (fst w1) (snd w1))) with beta.
+ unfold w1.
+ clear.
+ red; intros.
+ simpl in *.
+ apply (snd proof_weaken) with Gamma' ((beta,A1)::Delta'); intuition.
+ clear -X0.
+ unfold incl in *.
+ simpl in *.
+ intuition.
+ assert (H4: w1:A2||-).
+ apply (snd (IHA2 (fst w1) (snd w1))) with gamma.
+ unfold w1.
+ clear.
+ red; intros.
+ simpl in *.
+ apply (snd proof_weaken) with Gamma' ((gamma,A2)::Delta'); intuition.
+ clear -X0.
+ unfold incl in *.
+ simpl in *.
+ intuition.
+ destruct (H1 w1 H2 (H3,H4)) as [c Hc].
+ unfold w1 in Hc.
+ simpl in Hc.
+ exists (cm (mu beta (cm (mu gamma c) (mut y (cm (injr (tvar y)) (evar alpha))))) (mut y (cm (injl (tvar y)) (evar alpha)))).
+ (*eexists.*)
+ apply Cut with A1.
+ apply Mu.
+ apply Cut with A2.
+ apply Mu.
+ apply Hc.
+ apply H.
+ destruct H0.
+ intuition.
+ destruct H0; unfold incl in *; simpl in *; intuition.
+ apply MuT.
+ apply Cut with (Disj A1 A2).
+ apply DisjRR.
+ apply AxR.
+ intuition.
+ apply AxL.
+ intuition.
+ apply H.
+ destruct H0.
+ intuition.
+ destruct H0; unfold incl in *; simpl in *; intuition.
+ apply MuT.
+ apply Cut with (Disj A1 A2).
+ apply DisjRL.
+ apply AxR.
+ intuition.
+ apply AxL.
+ intuition.
+
+ (* Induction case - Conjunction *)
+ split.
+ split.
+ split.
+
+ intro H.
+ red in H.
+ change (forall w':K, (Gamma,Delta) <= w' -> refutes w' A1 + refutes w' A2
+ -> Kexploding w') in H.
+ set (w1 := (Gamma, (beta,A1)::Delta)).
+ assert (H11 : (Gamma,Delta) <= w1).
+ unfold w1.
+ clear;red;simpl;intuition.
+ assert (H12 : @refutes K w1 A1).
+ apply (snd (IHA1 (fst w1) (snd w1))) with beta.
+ unfold w1; simpl.
+ clear.
+ red.
+ intros.
+ apply (snd (proof_weaken)) with Gamma' ((beta,A1)::Delta').
+ assumption.
+ intuition.
+ red;simpl;intuition.
+ assert (H13:Kexploding w1).
+ intuition.
+ destruct H13 as [c1 Hc1].
+ unfold w1 in Hc1; simpl in Hc1.
+ set (w2 := (Gamma, (beta,A2)::Delta)).
+ assert (H21 : (Gamma,Delta) <= w2).
+ unfold w2.
+ clear;red;simpl;intuition.
+ assert (H22 : @refutes K w2 A2).
+ apply (snd (IHA2 (fst w2) (snd w2))) with beta.
+ unfold w2; simpl.
+ clear.
+ red.
+ intros.
+ apply (snd (proof_weaken)) with Gamma' ((beta,A2)::Delta').
+ assumption.
+ intuition.
+ red;simpl;intuition.
+ assert (H23:Kexploding w2).
+ intuition.
+ destruct H23 as [c2 Hc2].
+ unfold w2 in Hc2; simpl in Hc2.
+ clear - Hc1 Hc2.
+ eexists.
+ apply ConjR.
+ apply Mu.
+ apply Hc1.
+ apply Mu.
+ apply Hc2.
+
+ intros x H.
+ change (forall w':K, (Gamma,Delta) <= w' -> refutes w' A1 + refutes w' A2
+ -> Kexploding w').
+ intros w' H0 H1.
+ destruct H1.
+ destruct ((snd (fst (IHA1 (fst w') (snd w')))) r) as [e1 He1].
+ exists (cm (mu beta (cm (tvar x) (projl (evar beta)))) (e1)).
+ apply Cut with A1.
+ apply H.
+ destruct H0;intuition.
+ destruct H0;intuition.
+ apply Mu.
+ apply Cut with (Conj A1 A2).
+ apply AxR.
+ left;reflexivity.
+ apply ConjLL.
+ apply AxL.
+ left;reflexivity.
+ apply He1.
+ destruct ((snd (fst (IHA2 (fst w') (snd w')))) r) as [e2 He2].
+ exists (cm (mu beta (cm (tvar x) (projr (evar beta)))) (e2)).
+ apply Cut with A2.
+ apply H.
+ destruct H0;intuition.
+ destruct H0;intuition.
+ apply Mu.
+ apply Cut with (Conj A1 A2).
+ apply AxR.
+ left;reflexivity.
+ apply ConjLR.
+ apply AxL.
+ left;reflexivity.
+ apply He2.
+
+ intro H.
+ set (w' := ((y,Conj A1 A2)::Gamma, Delta)).
+ red in H.
+ simpl in H.
+ assert (H0 : (Gamma,Delta) <= w').
+ unfold w';clear;red;simpl;intuition.
+ apply H in H0.
+ destruct H0 as [c Hc].
+ unfold w' in Hc; simpl in Hc.
+ eexists.
+ apply MuT.
+ apply Hc.
+ change (forall w'':K, w' <= w'' -> refutes w'' A1 + refutes w'' A2
+ -> Kexploding w'').
+ intros w'' H1 H2.
+ destruct H2.
+ assert (H2 := (snd (fst (IHA1 (fst w'') (snd w'')))) r).
+ destruct H2 as [e He].
+ eexists.
+ apply Cut with (Conj A1 A2).
+ apply AxR.
+ apply H1.
+ left;reflexivity.
+ apply ConjLL.
+ apply He.
+ assert (H2 := (snd (fst (IHA2 (fst w'') (snd w'')))) r).
+ destruct H2 as [e He].
+ eexists.
+ apply Cut with (Conj A1 A2).
+ apply AxR.
+ apply H1.
+ left;reflexivity.
+ apply ConjLR.
+ apply He.
+
+ intros alpha H.
+ intros w' H0 H1.
+ simpl in *.
+ change (forall w'':K, w' <= w'' -> refutes w'' A1 + refutes w'' A2
+ -> Kexploding w'') in H1.
+ set (w1 := (fst w', (beta,A1)::snd w')).
+ assert (H11 : w' <= w1).
+ unfold w1;simpl;red;simpl;intuition.
+ assert (H12 : @refutes K w1 A1).
+ apply (snd (IHA1 (fst w1) (snd w1))) with beta.
+ unfold w1; simpl.
+ clear.
+ red.
+ intros.
+ apply (snd (proof_weaken)) with Gamma' ((beta,A1)::Delta').
+ assumption.
+ apply incl_refl.
+ clear -X0.
+ red in X0 |- *.
+ simpl in *; intuition.
+ set (w2 := (fst w', (beta,A2)::snd w')).
+ assert (H21 : w' <= w2).
+ unfold w2;simpl;red;simpl;intuition.
+ assert (H22 : @refutes K w2 A2).
+ apply (snd (IHA2 (fst w2) (snd w2))) with beta.
+ unfold w2; simpl.
+ clear.
+ red.
+ intros.
+ apply (snd (proof_weaken)) with Gamma' ((beta,A2)::Delta').
+ assumption.
+ apply incl_refl.
+ clear -X0.
+ red in X0 |- *.
+ simpl in *; intuition.
+ assert (Hc1 : Kexploding w1).
+ intuition.
+ assert (Hc2 : Kexploding w2).
+ intuition.
+ destruct Hc1 as [c1 Hc1].
+ destruct Hc2 as [c2 Hc2].
+ unfold w1 in Hc1; simpl in Hc1.
+ unfold w2 in Hc2; simpl in Hc2.
+ eexists.
+ apply Cut with (Conj A1 A2).
+ apply ConjR.
+ apply Mu.
+ apply Hc1.
+ apply Mu.
+ apply Hc2.
+ apply H.
+ destruct H0; intuition.
+ destruct H0; intuition.
+ apply AxL.
+ left; reflexivity.
+ Defined.
+End Context_Semantics.
+
+Lemma forces_cxt_id : forall Gamma Delta, @forces_cxt K (Gamma,Delta) Gamma.
+Proof.
+ induction Gamma.
+ simpl.
+ constructor.
+ simpl.
+ intro Delta.
+ split.
+ apply (snd (fst (fst (Kcompleteness (snd a) (a::Gamma) Delta)))) with (fst a).
+ red.
+ intros.
+ apply (snd (fst proof_weaken)) with ((fst a, snd a)::Gamma') Delta'.
+ assumption.
+ clear -X.
+ assert (In a Gamma').
+ intuition.
+ replace ((fst a, snd a)) with a.
+ red.
+ simpl.
+ intuition.
+ destruct a; auto.
+ apply incl_refl.
+ apply forces_cxt_mon with (Gamma, Delta).
+ auto.
+ simpl.
+ red.
+ simpl.
+ unfold incl.
+ intuition.
+Defined.
+
+Lemma refutes_cxt_id : forall Gamma Delta, @refutes_cxt K (Gamma,Delta) Delta.
+Proof.
+ induction Delta.
+ simpl.
+ constructor.
+ simpl.
+ split.
+ apply (snd (( (Kcompleteness (snd a) Gamma (a::Delta))))) with (fst a).
+ red.
+ intros.
+ apply (snd (proof_weaken)) with Gamma' ((fst a, snd a)::Delta').
+ assumption.
+ apply incl_refl.
+ clear -X0.
+ assert (In a Delta').
+ intuition.
+ replace ((fst a, snd a)) with a.
+ red.
+ simpl.
+ intuition.
+ destruct a; auto.
+ apply refutes_cxt_mon with (Gamma, Delta).
+ auto.
+ simpl.
+ red.
+ simpl.
+ unfold incl.
+ intuition.
+Defined.
+
+Definition NbE : forall c Gamma Delta, proof_cmd c Gamma Delta ->
+ {c':cmd & proof_cmd c' Gamma Delta}.
+Proof.
+ intros c Gamma Delta H.
+ destruct c as [t e].
+ inversion H.
+ clear -H2 H5.
+ assert (HGamma : @forces_cxt K (Gamma,Delta) Gamma).
+ apply forces_cxt_id.
+ assert (HDelta : @refutes_cxt K (Gamma,Delta) Delta).
+ apply refutes_cxt_id.
+ assert (H2' := (snd (fst (soundness K))) Gamma t A Delta H2 (Gamma,Delta) HGamma HDelta).
+ assert (H5' := (snd (soundness K)) Gamma e A Delta H5 (Gamma,Delta) HGamma HDelta).
+ unfold refutes in H5'.
+ simpl in H5'.
+ apply H5' in H2'.
+ red in H2'.
+ simpl in H2'.
+ exact H2'.
+ apply Kle_refl.
+Defined.
+
+(** Now some normalisation tests *)
+Parameter o : var_formula.
+Parameters x z : var_trm.
+
+(** Example 0: a simplest possible beta-reduction *)
+Definition cmd0 : cmd := cm (lam x (tvar x)) (app (tvar y) (evar beta)).
+Definition Gamma0 := ((y,Atom o)::nil).
+Definition Delta0 := ((beta,Atom o)::nil).
+Definition proof0 : proof_cmd cmd0 Gamma0 Delta0.
+Proof.
+ apply Cut with (Imp (Atom o) (Atom o)).
+ apply ImpR.
+ apply AxR.
+ left; reflexivity.
+ apply ImpL.
+ apply AxR.
+ left; reflexivity.
+ apply AxL.
+ left; reflexivity.
+Defined.
+Definition test0 := projT1 (NbE proof0).
+(* Time Eval vm_compute in test0. *)
+
+(** Example 1: a seemingly more complex beta-reduction, which reduces to the
+ same thing as example 0 *)
+Definition cmd1 : cmd := cm (lam x (tvar y)) (app (tvar x) (evar beta)).
+Definition Gamma1 := ((y,Atom o)::(x,Atom o)::nil).
+Definition Delta1 := ((beta,Atom o)::nil).
+Definition proof1 : proof_cmd cmd1 Gamma1 Delta1.
+Proof.
+ apply Cut with (Imp (Atom o) (Atom o)).
+ apply ImpR.
+ apply AxR.
+ right.
+ left; reflexivity.
+ apply ImpL.
+ apply AxR.
+ right.
+ left; reflexivity.
+ apply AxL.
+ left; reflexivity.
+Defined.
+Definition test1 := projT1 (NbE proof1).
+(* Time Eval vm_compute in test1. *)
+
+(** Example 2: The critical pair indicates we are indeed in
+ call-by-name. To check that the CBN behaviour does not depend on the choice of a proof,
+ we try it with two different proofs. *)
+Definition c' : cmd := cm (tvar x) (evar gamma).
+Definition c'' : cmd := cm (tvar z) (evar delta).
+Definition cmd2 := cm (mu beta c') (mut y c'').
+Definition proof2 : proof_cmd cmd2 ((x,Atom o)::(z,Atom o)::nil) ((gamma,Atom o)::(delta,Atom o)::nil).
+Proof.
+ apply Cut with (Atom o).
+ apply Mu.
+ apply Cut with (Atom o).
+ apply AxR.
+ left;reflexivity.
+ apply AxL.
+ right;left;reflexivity.
+ apply MuT.
+ apply Cut with (Atom o).
+ apply AxR.
+ right;right;left;reflexivity.
+ apply AxL.
+ right;left;reflexivity.
+Defined.
+Definition proof2' : proof_cmd cmd2 ((z,Atom o)::(x,Atom o)::nil) ((delta,Atom o)::(gamma,Atom o)::nil).
+Proof.
+ apply Cut with (Atom o).
+ apply Mu.
+ apply Cut with (Atom o).
+ apply AxR.
+ right;left;reflexivity.
+ apply AxL.
+ right;right;left;reflexivity.
+ apply MuT.
+ apply Cut with (Atom o).
+ apply AxR.
+ right;left;reflexivity.
+ apply AxL.
+ left;reflexivity.
+Defined.
+Definition test2 := projT1 (NbE proof2).
+Definition test2' := projT1 (NbE proof2').
+(* Time Eval vm_compute in test2. *)
+(* Time Eval vm_compute in test2'. *)
+
+(** Example 3: verifying reduction for implication *)
+Definition cmd3 := cm (lam x (tvar z)) (app (tvar y) (evar gamma)).
+Definition proof3 : proof_cmd cmd3 ((y,Atom o)::(z,Atom o)::nil) ((gamma,Atom o)::nil).
+Proof.
+ apply Cut with (Imp (Atom o) (Atom o)).
+ apply ImpR.
+ apply AxR.
+ right;right;left;reflexivity.
+ apply ImpL.
+ apply AxR.
+ left;reflexivity.
+ apply AxL.
+ left;reflexivity.
+Defined.
+Definition test3 := projT1 (NbE proof3).
+(* Time Eval vm_compute in test3. *)
+
+(** Example 4: verifying reduction for mu *)
+Definition cmd4 := cm (mu gamma (cm (tvar y) (evar delta))) (evar beta).
+Definition proof4 : proof_cmd cmd4 ((y,Atom o)::nil) ((beta,Atom o)::(delta,Atom o)::nil).
+Proof.
+ apply Cut with (Atom o).
+ apply Mu.
+ apply Cut with (Atom o).
+ apply AxR.
+ left;reflexivity.
+ apply AxL.
+ right;right;left;reflexivity.
+ apply AxL.
+ left;reflexivity.
+Defined.
+Definition test4 := projT1 (NbE proof4).
+(* Time Eval vm_compute in test4. *)
+
+(** Example 4': verifying reduction for mu *)
+Definition cmd4' := cm (mu gamma (cm (tvar y) (evar gamma))) (evar beta).
+Definition proof4' : proof_cmd cmd4' ((y,Atom o)::nil) ((beta,Atom o)::nil).
+Proof.
+ apply Cut with (Atom o).
+ apply Mu.
+ apply Cut with (Atom o).
+ apply AxR.
+ left;reflexivity.
+ apply AxL.
+ left;reflexivity.
+ apply AxL.
+ left;reflexivity.
+Defined.
+Definition test4' := projT1 (NbE proof4').
+(* Time Eval vm_compute in test4'. *)
+
+(** Example 5: verifying reduction for conjunction 1 *)
+Definition cmd5 := cm (paire (tvar x) (tvar z)) (projl (evar beta)).
+Definition proof5 : proof_cmd cmd5 ((x,Atom o)::(z,Atom o)::nil) ((beta,Atom o)::nil).
+Proof.
+ apply Cut with (Conj (Atom o) (Atom o)).
+ apply ConjR.
+ apply AxR.
+ left;reflexivity.
+ apply AxR.
+ right;left;reflexivity.
+ apply ConjLL.
+ apply AxL.
+ left;reflexivity.
+Defined.
+Definition test5 := projT1 (NbE proof5).
+(* Time Eval vm_compute in test5. *)
+
+(** Example 6: verifying reduction for conjunction 2 *)
+Definition cmd6 := cm (paire (tvar x) (tvar z)) (projr (evar beta)).
+Definition proof6 : proof_cmd cmd6 ((x,Atom o)::(z,Atom o)::nil) ((beta,Atom o)::nil).
+Proof.
+ apply Cut with (Conj (Atom o) (Atom o)).
+ apply ConjR.
+ apply AxR.
+ left;reflexivity.
+ apply AxR.
+ right;left;reflexivity.
+ apply ConjLR.
+ apply AxL.
+ left;reflexivity.
+Defined.
+Definition test6 := projT1 (NbE proof6).
+(* Time Eval vm_compute in test6. *)
+
+(** Example 7: verifying reduction for disjunction 1 *)
+Definition cmd7 := cm (injl (tvar z)) (disj (evar gamma) (evar delta)).
+Definition proof7 : proof_cmd cmd7 ((z,Atom o)::nil) ((gamma,Atom o)::(delta,Atom o)::nil).
+Proof.
+ apply Cut with (Disj (Atom o) (Atom o)).
+ apply DisjRL.
+ apply AxR.
+ left;reflexivity.
+ apply DisjL.
+ apply AxL.
+ left;reflexivity.
+ apply AxL.
+ right;left;reflexivity.
+Defined.
+Definition test7 := projT1 (NbE proof7).
+(* Time Eval vm_compute in test7. *)
+
+(** Example 8: verifying reduction for disjunction 2 *)
+Definition cmd8 := cm (injr (tvar z)) (disj (evar gamma) (evar delta)).
+Definition proof8 : proof_cmd cmd8 ((z,Atom o)::nil) ((gamma,Atom o)::(delta,Atom o)::nil).
+Proof.
+ apply Cut with (Disj (Atom o) (Atom o)).
+ apply DisjRR.
+ apply AxR.
+ left;reflexivity.
+ apply DisjL.
+ apply AxL.
+ left;reflexivity.
+ apply AxL.
+ right;left;reflexivity.
+Defined.
+Definition test8 := projT1 (NbE proof8).
+(* Time Eval vm_compute in test8. *)
+
diff --git a/kripke_completeness/class_comp_cbv.v b/kripke_completeness/class_comp_cbv.v
new file mode 100644
index 0000000..6cda589
--- /dev/null
+++ b/kripke_completeness/class_comp_cbv.v
@@ -0,0 +1,3463 @@
+(** Formalised proof of completeness of full Classical predicate logic
+ with respect to Kripke-style models.
+
+ Danko Ilik, January 2010
+
+ Normalisation examples at end of file.
+
+ Note: Soundness theorem quantifier cases need to be finished.
+*)
+Require Import stdlib_Type.
+Require Import EqNat.
+Require Import Peano_dec.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+(** printing <= $\le$ #≤# *)
+(** printing ||-- $\Vdash$ #⊩# *)
+(** printing ||- $\Vdash$ #⊩# *)
+(** printing ||+ $\Vdash_s$ #⊩s # *)
+
+Open Scope type_scope.
+
+Definition var_trm := nat.
+Definition var_ect := nat.
+
+(** Commands, terms, and evaluation contexts *)
+Inductive cmd : Set :=
+| cm : trm -> ect -> cmd
+with trm : Set :=
+| tvar : var_trm -> trm
+| lam : var_trm -> trm -> trm
+| injl : trm -> trm
+| injr : trm -> trm
+| paire : trm -> trm -> trm
+| mu : var_ect -> cmd -> trm
+| tall : trm -> trm
+| tex : trm -> trm
+with ect : Set :=
+| evar : var_ect -> ect
+| app : trm -> ect -> ect
+| disj : ect -> ect -> ect
+| projl : ect -> ect
+| projr : ect -> ect
+| mut : var_trm -> cmd -> ect
+| eall : ect -> ect
+| eex : ect -> ect.
+
+(** Abstract (decidable) sets of fresh variables, and predicate and function symbols *)
+Parameter var_free : Set.
+Parameter var_free_dec : forall x y:var_free, {x=y}+{~x=y}.
+Parameters predicate function : Set.
+Notation "x \notin L" := (In x L -> Empty_set) (at level 69).
+Hypothesis fresh_fvar : forall L, sigT (fun x:var_free => x \notin L).
+
+(** Individual terms. [bvar]s are deBruijn indices which represent bound variables, while [fvar]s are names that represent free variables *)
+Inductive indiv : Set :=
+ bvar : nat -> indiv
+| fvar : var_free -> indiv
+| func : function -> list indiv -> indiv.
+
+Inductive typ : Set :=
+ Atom : predicate -> list indiv -> typ
+| Imp : typ -> typ -> typ
+| Disj : typ -> typ -> typ
+| Conj : typ -> typ -> typ
+| All : typ -> typ
+| Ex : typ -> typ.
+
+(** Substituting at de Bruijn variable k *)
+
+Fixpoint subst_indiv (k:nat)(a:indiv)(d:indiv) {struct a} : indiv :=
+ let subst_list := fix subst_list (l:list indiv) {struct l} : list indiv :=
+ match l with
+ | nil => nil
+ | cons a' l' => cons (subst_indiv k a' d) (subst_list l')
+ end in
+ match a with
+ | func f ld =>
+ func f (subst_list ld)
+ | fvar x => fvar x
+ | bvar i => if beq_nat k i then d else bvar i
+ end.
+
+Fixpoint subst (k:nat)(A:typ)(d:indiv) {struct A} : typ :=
+ let subst_list := fix subst_list (l:list indiv) {struct l} : list indiv :=
+ match l with
+ | nil => nil
+ | cons a' l' => cons (subst_indiv k a' d) (subst_list l')
+ end in
+ match A with
+ | Ex A1 => Ex (subst (S k) A1 d)
+ | All A1 => All (subst (S k) A1 d)
+ | Conj A1 A2 => Conj (subst k A1 d) (subst k A2 d)
+ | Disj A1 A2 => Disj (subst k A1 d) (subst k A2 d)
+ | Imp A1 A2 => Imp (subst k A1 d) (subst k A2 d)
+ | Atom P ld =>
+ Atom P (subst_list ld)
+ end.
+
+Notation "A ^^ d" := (subst 0 A d) (at level 67).
+Notation "A ^ x" := (subst 0 A (fvar x)).
+
+(** A formula is [locl] ("locally closed") when all [bvar]s are bound
+ by quantifiers (while there might be [fvar]s around) *)
+Definition locl (A:typ) := forall k d, (subst k A d) = A.
+Definition locli (d':indiv) := forall k d, (subst_indiv k d' d) = d'.
+
+Let cxt_trm := list (var_trm*typ).
+Let cxt_ect := list (var_ect*typ).
+
+Reserved Notation "c : ( Gamma |- Delta ) " (at level 70).
+Reserved Notation "Gamma |- ( t : A ) ; Delta"
+ (at level 70, A at next level, t at next level).
+Reserved Notation "Gamma ; ( e : A ) |- Delta"
+ (at level 70, A at next level, e at next level).
+
+(** The LK-mu-mu-tilde sequent calculus with cofinite quantification *)
+Inductive proof_cmd : cmd -> cxt_trm -> cxt_ect -> Set :=
+| Cut v e Gamma Delta A :
+ proof_trm Gamma v A Delta ->
+ proof_ect Gamma e A Delta ->
+ proof_cmd (cm v e) Gamma Delta
+
+where "c : ( Gamma |- Delta )" := (proof_cmd c Gamma Delta)
+
+with proof_trm : cxt_trm -> trm -> typ -> cxt_ect -> Set :=
+| AxR Gamma a A Delta :
+ In (a,A) Gamma ->
+ proof_trm Gamma (tvar a) A Delta
+
+| Mu Gamma alpha c A Delta :
+ proof_cmd c Gamma ((alpha,A)::Delta) ->
+ proof_trm Gamma (mu alpha c) A Delta
+
+| ImpR Gamma a t A B Delta :
+ proof_trm ((a,A)::Gamma) t B Delta ->
+ proof_trm Gamma (lam a t) (Imp A B) Delta
+
+| DisjRL Gamma v A1 A2 Delta :
+ proof_trm Gamma v A1 Delta ->
+ proof_trm Gamma (injl v) (Disj A1 A2) Delta
+
+| DisjRR Gamma v A1 A2 Delta :
+ proof_trm Gamma v A2 Delta ->
+ proof_trm Gamma (injr v) (Disj A1 A2) Delta
+
+| ConjR Gamma v1 v2 A1 A2 Delta :
+ proof_trm Gamma v1 A1 Delta ->
+ proof_trm Gamma v2 A2 Delta ->
+ proof_trm Gamma (paire v1 v2) (Conj A1 A2) Delta
+
+| AllR Gamma t A Delta L :
+ locl (All A) ->
+ (forall x, x \notin L -> proof_trm Gamma t (A^x) Delta) ->
+ proof_trm Gamma (tall t) (All A) Delta
+
+| ExR Gamma t A Delta d : locli d ->
+ proof_trm Gamma t (A^^d) Delta ->
+ proof_trm Gamma (tex t) (Ex A) Delta
+
+where "Gamma |- ( t : A ) ; Delta" := (proof_trm Gamma t A Delta)
+
+with proof_ect : cxt_trm -> ect -> typ -> cxt_ect -> Set :=
+| AxL Gamma alpha A Delta :
+ In (alpha,A) Delta ->
+ proof_ect Gamma (evar alpha) A Delta
+
+| MuT Gamma a c A Delta :
+ proof_cmd c ((a,A)::Gamma) Delta ->
+ proof_ect Gamma (mut a c) A Delta
+
+| ImpL Gamma v e A B Delta :
+ proof_trm Gamma v A Delta ->
+ proof_ect Gamma e B Delta ->
+ proof_ect Gamma (app v e) (Imp A B) Delta
+
+| DisjL Gamma e1 e2 A1 A2 Delta :
+ proof_ect Gamma e1 A1 Delta ->
+ proof_ect Gamma e2 A2 Delta ->
+ proof_ect Gamma (disj e1 e2) (Disj A1 A2) Delta
+
+| ConjLL Gamma e A1 A2 Delta :
+ proof_ect Gamma e A1 Delta ->
+ proof_ect Gamma (projl e) (Conj A1 A2) Delta
+
+| ConjLR Gamma e A1 A2 Delta :
+ proof_ect Gamma e A2 Delta ->
+ proof_ect Gamma (projr e) (Conj A1 A2) Delta
+
+| AllL Gamma e A Delta d : locli d ->
+ proof_ect Gamma e (A^^d) Delta ->
+ proof_ect Gamma (eall e) (All A) Delta
+
+| ExL Gamma e A Delta L :
+ locl (Ex A) ->
+ (forall x, x \notin L -> proof_ect Gamma e (A^x) Delta) ->
+ proof_ect Gamma (eex e) (Ex A) Delta
+
+where "Gamma ; ( e : A ) |- Delta" := (proof_ect Gamma e A Delta)
+.
+
+(* (* Combined Scheme only works for sort Prop, so I have to give the *) *)
+(* (* definition by hand by adapting a copy-paste of the Prop version: *) *)
+(* Scheme proof_cmd_rect' := Induction for proof_cmd Sort Prop *)
+(* with proof_trm_rect' := Induction for proof_trm Sort Prop *)
+(* with proof_ect_rect' := Induction for proof_ect Sort Prop. *)
+(* Combined Scheme proof_rect'' from proof_cmd_rect', proof_trm_rect', *)
+(* proof_ect_rect'. *)
+(* Unset Printing Notations. *)
+(* Print proof_rect''. *)
+(* Set Printing Notations. *)
+
+Scheme proof_cmd_rect' := Induction for proof_cmd Sort Type
+with proof_trm_rect' := Induction for proof_trm Sort Type
+with proof_ect_rect' := Induction for proof_ect Sort Type.
+
+(** We have to make the mutual-induction predicate "by hand" since the Combined Scheme Coq command works only in sort Prop, not in Type. *)
+Definition proof_rect' := fun
+ (P : forall (c : cmd) (c0 : cxt_trm) (c1 : cxt_ect),
+ proof_cmd c c0 c1 -> Type)
+ (P0 : forall (c : cxt_trm) (t : trm) (t0 : typ) (c0 : cxt_ect),
+ proof_trm c t t0 c0 -> Type)
+ (P1 : forall (c : cxt_trm) (e : ect) (t : typ) (c0 : cxt_ect),
+ proof_ect c e t c0 -> Type)
+ (f : forall (v : trm) (e : ect) (Gamma : cxt_trm)
+ (Delta : cxt_ect) (A : typ) (p : proof_trm Gamma v A Delta),
+ P0 Gamma v A Delta p ->
+ forall p0 : proof_ect Gamma e A Delta,
+ P1 Gamma e A Delta p0 -> P (cm v e) Gamma Delta (Cut p p0))
+ (f0 : forall (Gamma : list (prod var_trm typ)) (a : var_trm)
+ (A : typ) (Delta : cxt_ect) (i : In (pair a A) Gamma),
+ P0 Gamma (tvar a) A Delta (AxR (Gamma:=Gamma) (a:=a) (A:=A) Delta i))
+ (f1 : forall (Gamma : cxt_trm) (alpha : var_ect)
+ (c : cmd) (A : typ) (Delta : list (prod var_ect typ))
+ (p : proof_cmd c Gamma (cons (pair alpha A) Delta)),
+ P c Gamma (cons (pair alpha A) Delta) p ->
+ P0 Gamma (mu alpha c) A Delta (Mu p))
+ (f2 : forall (Gamma : list (prod var_trm typ)) (a : var_trm)
+ (t : trm) (A B : typ) (Delta : cxt_ect)
+ (p : proof_trm (cons (pair a A) Gamma) t B Delta),
+ P0 (cons (pair a A) Gamma) t B Delta p ->
+ P0 Gamma (lam a t) (Imp A B) Delta (ImpR p))
+ (f3 : forall (Gamma : cxt_trm) (v : trm) (A1 A2 : typ)
+ (Delta : cxt_ect) (p : proof_trm Gamma v A1 Delta),
+ P0 Gamma v A1 Delta p ->
+ P0 Gamma (injl v) (Disj A1 A2) Delta (DisjRL A2 p))
+ (f4 : forall (Gamma : cxt_trm) (v : trm) (A1 A2 : typ)
+ (Delta : cxt_ect) (p : proof_trm Gamma v A2 Delta),
+ P0 Gamma v A2 Delta p ->
+ P0 Gamma (injr v) (Disj A1 A2) Delta (DisjRR A1 p))
+ (f5 : forall (Gamma : cxt_trm) (v1 v2 : trm) (A1 A2 : typ)
+ (Delta : cxt_ect) (p : proof_trm Gamma v1 A1 Delta),
+ P0 Gamma v1 A1 Delta p ->
+ forall p0 : proof_trm Gamma v2 A2 Delta,
+ P0 Gamma v2 A2 Delta p0 ->
+ P0 Gamma (paire v1 v2) (Conj A1 A2) Delta (ConjR p p0))
+ (f6 : forall (Gamma : cxt_trm) (t : trm) (A : typ)
+ (Delta : cxt_ect) (L : list var_free) (l : locl (All A))
+ (p : forall x : var_free,
+ (In x L -> Empty_set) ->
+ proof_trm Gamma t (subst O A (fvar x)) Delta),
+ (forall (x : var_free) (e : In x L -> Empty_set),
+ P0 Gamma t (subst O A (fvar x)) Delta (p x e)) ->
+ P0 Gamma (tall t) (All A) Delta (AllR (A:=A) (L:=L) l p))
+ (f7 : forall (Gamma : cxt_trm) (t : trm) (A : typ)
+ (Delta : cxt_ect) (d : indiv) (l : locli d)
+ (p : proof_trm Gamma t (subst O A d) Delta),
+ P0 Gamma t (subst O A d) Delta p ->
+ P0 Gamma (tex t) (Ex A) Delta (ExR (A:=A) (d:=d) l p))
+ (f8 : forall (Gamma : cxt_trm) (alpha : var_ect)
+ (A : typ) (Delta : list (prod var_ect typ))
+ (i : In (pair alpha A) Delta),
+ P1 Gamma (evar alpha) A Delta
+ (AxL Gamma (alpha:=alpha) (A:=A) (Delta:=Delta) i))
+ (f9 : forall (Gamma : list (prod var_trm typ)) (a : var_trm)
+ (c : cmd) (A : typ) (Delta : cxt_ect)
+ (p : proof_cmd c (cons (pair a A) Gamma) Delta),
+ P c (cons (pair a A) Gamma) Delta p ->
+ P1 Gamma (mut a c) A Delta (MuT p))
+ (f10 : forall (Gamma : cxt_trm) (v : trm) (e : ect)
+ (A B : typ) (Delta : cxt_ect) (p : proof_trm Gamma v A Delta),
+ P0 Gamma v A Delta p ->
+ forall p0 : proof_ect Gamma e B Delta,
+ P1 Gamma e B Delta p0 ->
+ P1 Gamma (app v e) (Imp A B) Delta (ImpL p p0))
+ (f11 : forall (Gamma : cxt_trm) (e1 e2 : ect) (A1 A2 : typ)
+ (Delta : cxt_ect) (p : proof_ect Gamma e1 A1 Delta),
+ P1 Gamma e1 A1 Delta p ->
+ forall p0 : proof_ect Gamma e2 A2 Delta,
+ P1 Gamma e2 A2 Delta p0 ->
+ P1 Gamma (disj e1 e2) (Disj A1 A2) Delta (DisjL p p0))
+ (f12 : forall (Gamma : cxt_trm) (e : ect) (A1 A2 : typ)
+ (Delta : cxt_ect) (p : proof_ect Gamma e A1 Delta),
+ P1 Gamma e A1 Delta p ->
+ P1 Gamma (projl e) (Conj A1 A2) Delta (ConjLL A2 p))
+ (f13 : forall (Gamma : cxt_trm) (e : ect) (A1 A2 : typ)
+ (Delta : cxt_ect) (p : proof_ect Gamma e A2 Delta),
+ P1 Gamma e A2 Delta p ->
+ P1 Gamma (projr e) (Conj A1 A2) Delta (ConjLR A1 p))
+ (f14 : forall (Gamma : cxt_trm) (e : ect) (A : typ)
+ (Delta : cxt_ect) (d : indiv) (l : locli d)
+ (p : proof_ect Gamma e (subst O A d) Delta),
+ P1 Gamma e (subst O A d) Delta p ->
+ P1 Gamma (eall e) (All A) Delta (AllL (A:=A) (d:=d) l p))
+ (f15 : forall (Gamma : cxt_trm) (e : ect) (A : typ)
+ (Delta : cxt_ect) (L : list var_free) (l : locl (Ex A))
+ (p : forall x : var_free,
+ (In x L -> Empty_set) ->
+ proof_ect Gamma e (subst O A (fvar x)) Delta),
+ (forall (x : var_free) (e0 : In x L -> Empty_set),
+ P1 Gamma e (subst O A (fvar x)) Delta (p x e0)) ->
+ P1 Gamma (eex e) (Ex A) Delta (ExL (A:=A) (L:=L) l p)) =>
+ pair
+ (pair
+ (proof_cmd_rect' (P:=P) (P0:=P0) (P1:=P1) f f0 f1 f2 f3 f4 f5 f6 f7 f8 f9
+ f10 f11 f12 f13 f14 f15)
+ (proof_trm_rect' (P:=P) (P0:=P0) (P1:=P1) f f0 f1 f2 f3 f4 f5 f6 f7 f8
+ f9 f10 f11 f12 f13 f14 f15))
+ (proof_ect_rect' (P:=P) (P0:=P0) (P1:=P1) f f0 f1 f2 f3 f4 f5 f6 f7 f8
+ f9 f10 f11 f12 f13 f14 f15).
+
+Section Abstract_Semantics.
+ (** An abstract Kripke-style structure: [wle] is the preorder, [X] is
+ the exploding-world predicate, [aforces] is strong forcing of
+ atomic formulae. *)
+ Record Kripke : Type := {
+ world :> Set;
+ wle : world -> world -> Type;
+ wle_refl : forall w, wle w w;
+ wle_trans : forall w w' w'', wle w w' -> wle w' w'' -> wle w w'';
+ exploding : world -> Set;
+ aforces : world -> predicate -> list indiv -> Set;
+ aforces_mon : forall w P ld, @aforces w P ld ->
+ forall w', wle w w' -> @aforces w' P ld
+ }.
+ Notation "w <= w'" := (wle w w').
+
+ Variable K:Kripke.
+
+ (** The continuations monad we will use to extend forcing to
+ composite formulas *)
+ Definition Kont (F:K->typ->Type)(w:K)(A:typ) :=
+ forall w1, w <= w1 ->
+ (forall w2, w1 <= w2 -> F w2 A -> exploding w2) -> exploding w1.
+ Hint Unfold Kont.
+
+ (** Strong forcing extended to composite formulas. The expression
+ [sforces w A] stands for strong forcing, while [Kont sforces w A]
+ stands for (weak) forcing of A in the world w. *)
+ Fixpoint sforces' (n:nat)(w:K)(A:typ) {struct n} : Type :=
+ match n with
+ | O => Empty_set
+ | S m =>
+ match A with
+ | Atom P ld => aforces w P ld
+ | Imp A1 A2 => forall w', w <= w'-> Kont (sforces' m) w' A1 -> Kont (sforces' m) w' A2
+ | Disj A1 A2 => Kont (sforces' m) w A1 + Kont (sforces' m) w A2
+ | Conj A1 A2 => Kont (sforces' m) w A1 * Kont (sforces' m) w A2
+ | All A1 => forall w', w <= w' -> forall d, locli d -> Kont (sforces' m) w' (A1^^d)
+ | Ex A1 => sigT (fun d => (locli d) * Kont (sforces' m) w (A1^^d))
+ end
+ end.
+
+ Fixpoint depth (A:typ) : nat :=
+ match A with
+ | Atom _ _ => 1
+ | Imp A1 A2 => S (max (depth A1) (depth A2))
+ | Disj A1 A2 => S (max (depth A1) (depth A2))
+ | Conj A1 A2 => S (max (depth A1) (depth A2))
+ | All A1 => S (depth A1)
+ | Ex A1 => S (depth A1)
+ end.
+
+ (** However, we can not define strong forcing as simply as when we had no quantifiers, because we need to do a recursion on the _complexity_ of a formula, not just its structure; we do that using the measure of [depth] above. *)
+ Definition sforces (w:K)(A:typ) := sforces' (depth A) w A.
+ Hint Unfold sforces.
+
+ (** And now this section proves that the definition using the measure is correct *)
+ Section sforces_correctness.
+ Lemma depth_subst : forall A d k, depth (subst k A d) = depth A.
+ Proof.
+ induction A; simpl; intros.
+
+ reflexivity.
+
+ rewrite <- IHA1 with d k.
+ rewrite <- IHA2 with d k.
+ reflexivity.
+
+ rewrite <- IHA1 with d k.
+ rewrite <- IHA2 with d k.
+ reflexivity.
+
+ rewrite <- IHA1 with d k.
+ rewrite <- IHA2 with d k.
+ reflexivity.
+
+ rewrite <- IHA with d (S k).
+ reflexivity.
+
+ rewrite <- IHA with d (S k).
+ reflexivity.
+ Defined.
+
+ Lemma sforces'_S : forall n A w, le (depth A) n ->
+ (sforces' n w A <-> sforces' (S n) w A).
+ Proof.
+ induction n;
+ intros A w Hle.
+
+ inversion Hle.
+ destruct A;
+ inversion H0.
+
+ split.
+
+ (* Direction ==> *)
+ intro H.
+ simpl in H.
+ hnf.
+ assert (IHn1 := fun A w Hle => fst (IHn A w Hle)).
+ assert (IHn2 := fun A w Hle => snd (IHn A w Hle)).
+ clear IHn.
+ destruct A.
+
+ assumption.
+
+ intros w' ww' HKA1Sn.
+ intros w1 w'w1 k.
+ apply H with w'.
+ assumption.
+ intros w2 w1w2 HA1n.
+ apply HKA1Sn.
+ assumption.
+ intros w3 w2w3 HA1Sn.
+ apply HA1n.
+ assumption.
+ apply IHn2.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_l.
+ apply Hle.
+ assumption.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn1.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_r.
+ apply Hle.
+ assumption.
+
+ case H; intro H'.
+ left.
+ intros w1 ww1 k.
+ apply H'.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn1; auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_l.
+ apply Hle.
+ right.
+ intros w1 ww1 k.
+ apply H'.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn1; auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_r.
+ apply Hle.
+
+ destruct H as [H1 H2].
+ split.
+ intros w1 ww1 k.
+ apply H1.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn1;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_l.
+ apply Hle.
+ intros w1 ww1 k.
+ apply H2.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn1;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_r.
+ apply Hle.
+
+ intros.
+ intros w1 w'w1 k.
+ apply H with w1 d.
+ eauto using wle_trans.
+ assumption.
+ apply wle_refl.
+ intros w2 w1w2 HA.
+ apply k.
+ assumption.
+ apply IHn1;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ rewrite depth_subst.
+ assumption.
+
+ destruct H as [d [H1 H2]].
+ exists d.
+ split; [assumption | ].
+ intros w1 ww1 k.
+ apply H2.
+ assumption.
+ intros w2 w1w2 H3.
+ apply k.
+ assumption.
+ apply IHn1;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ rewrite depth_subst.
+ assumption.
+
+
+ (* Direction <== (exact copy-past of the other direction except for
+ swapping IHn1 <~> IHn2) *)
+ intro H.
+ simpl.
+ hnf in H.
+ assert (IHn1 := fun A w Hle => fst (IHn A w Hle)).
+ assert (IHn2 := fun A w Hle => snd (IHn A w Hle)).
+ clear IHn.
+ destruct A.
+
+ assumption.
+
+ intros w' ww' HKA1Sn.
+ intros w1 w'w1 k.
+ apply H with w'.
+ assumption.
+ intros w2 w1w2 HA1n.
+ apply HKA1Sn.
+ assumption.
+ intros w3 w2w3 HA1Sn.
+ apply HA1n.
+ assumption.
+ apply IHn1.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_l.
+ apply Hle.
+ assumption.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn2.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_r.
+ apply Hle.
+ assumption.
+
+ case H; intro H'.
+ left.
+ intros w1 ww1 k.
+ apply H'.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn2; auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_l.
+ apply Hle.
+ right.
+ intros w1 ww1 k.
+ apply H'.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn2; auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_r.
+ apply Hle.
+
+ destruct H as [H1 H2].
+ split.
+ intros w1 ww1 k.
+ apply H1.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn2;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_l.
+ apply Hle.
+ intros w1 ww1 k.
+ apply H2.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn2;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_r.
+ apply Hle.
+
+ intros.
+ intros w1 w'w1 k.
+ apply H with w1 d.
+ eauto using wle_trans.
+ assumption.
+ apply wle_refl.
+ intros w2 w1w2 HA.
+ apply k.
+ assumption.
+ apply IHn2;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ rewrite depth_subst.
+ assumption.
+
+ destruct H as [d [H1 H2]].
+ exists d.
+ split; [assumption | ].
+ intros w1 ww1 k.
+ apply H2.
+ assumption.
+ intros w2 w1w2 H3.
+ apply k.
+ assumption.
+ apply IHn2;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ rewrite depth_subst.
+ assumption.
+ Defined.
+
+ Definition sforces'_S1 := fun n A w Hle => fst (@sforces'_S n A w Hle).
+ Definition sforces'_S2 := fun n A w Hle => snd (@sforces'_S n A w Hle).
+
+ Lemma sforces'_mon1 : forall n m, le n m ->
+ forall A, le (depth A) n -> forall w, sforces' n w A -> sforces' m w A.
+ Proof.
+ intros n m Hle.
+ induction Hle.
+
+ intros.
+ auto.
+
+ intros.
+ apply sforces'_S1.
+ eauto using le_trans.
+ auto.
+ Defined.
+
+ Lemma sforces'_mon2 : forall n m, le n m ->
+ forall A, le (depth A) n -> forall w, sforces' m w A -> sforces' n w A.
+ Proof.
+ intros n m Hle.
+ induction Hle.
+
+ intros.
+ auto.
+
+ intros.
+ apply IHHle.
+ assumption.
+ apply sforces'_S2.
+ eauto using le_trans.
+ assumption.
+ Defined.
+
+ Lemma Kont_invar : forall F1 F2:K->typ->Type, forall A,
+ (forall w, F1 w A -> F2 w A) -> forall w, Kont F1 w A -> Kont F2 w A.
+ Proof.
+ intros F1 F2 A HF1F2.
+ intros w HF1.
+ intros w1 ww1 k1.
+ apply HF1.
+ assumption.
+ intros w2 w1w2 HF1'.
+ apply k1.
+ assumption.
+ auto.
+ Defined.
+
+ Lemma sforces_correct_Atom : forall w P ld,
+ sforces w (@Atom P ld) -> aforces w P ld.
+ Proof.
+ unfold sforces.
+ simpl.
+ auto.
+ Defined.
+
+ Lemma sforces_correct_Atom' : forall w P ld,
+ aforces w P ld -> sforces w (@Atom P ld).
+ Proof.
+ unfold sforces.
+ simpl.
+ auto.
+ Defined.
+
+ Lemma sforces_correct_Disj : forall w A B, sforces w (Disj A B) ->
+ Kont sforces w A + Kont sforces w B.
+ Proof.
+ intros w A B H.
+ unfold sforces in H.
+ simpl in H.
+ destruct H.
+
+ left.
+ change (Kont (sforces' (depth A)) w A).
+ generalize dependent k.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon2.
+ apply le_max_l.
+ constructor.
+
+ right.
+ change (Kont (sforces' (depth B)) w B).
+ generalize dependent k.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon2.
+ apply le_max_r.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_Disj' : forall w A B,
+ Kont sforces w A + Kont sforces w B -> sforces w (Disj A B).
+ Proof.
+ intros w A B H.
+ unfold sforces.
+ simpl.
+ destruct H.
+
+ left.
+ generalize dependent k.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon1.
+ apply le_max_l.
+ constructor.
+
+ right.
+ generalize dependent k.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon1.
+ apply le_max_r.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_Conj : forall w A B, sforces w (Conj A B) ->
+ Kont sforces w A * Kont sforces w B.
+ Proof.
+ intros w A B H.
+ unfold sforces in H.
+ simpl in H.
+ destruct H as [H1 H2].
+ split.
+
+ change (Kont (sforces' (depth A)) w A).
+ generalize dependent H1.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon2.
+ apply le_max_l.
+ constructor.
+
+ change (Kont (sforces' (depth B)) w B).
+ generalize dependent H2.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon2.
+ apply le_max_r.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_Conj' : forall w A B,
+ Kont sforces w A * Kont sforces w B -> sforces w (Conj A B).
+ Proof.
+ intros w A B H.
+ unfold sforces.
+ simpl.
+ destruct H as [H1 H2].
+ split.
+
+ generalize dependent H1.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon1.
+ apply le_max_l.
+ constructor.
+
+ generalize dependent H2.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon1.
+ apply le_max_r.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_Imp : forall w A B, sforces w (Imp A B) ->
+ forall w', w <= w' -> Kont sforces w' A -> Kont sforces w' B.
+ Proof.
+ intros w A B H.
+ unfold sforces in H.
+ simpl in H.
+ intros w' ww' HA.
+
+ change (Kont (sforces' (depth A)) w' A) in HA.
+ change (Kont (sforces' (depth B)) w' B).
+
+ apply Kont_invar with (sforces' (max (depth A) (depth B))).
+ apply sforces'_mon2.
+ apply le_max_r.
+ constructor.
+ apply H.
+ assumption.
+ generalize dependent HA.
+ apply Kont_invar.
+ apply sforces'_mon1.
+ apply le_max_l.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_Imp' : forall w A B,
+ (forall w', w <= w' -> Kont sforces w' A -> Kont sforces w' B)
+ -> sforces w (Imp A B).
+ Proof.
+ intros w A B H.
+ unfold sforces.
+ simpl.
+ intros w' ww' HA.
+
+ apply Kont_invar with (sforces' (depth B)).
+ intros w''.
+ apply sforces'_mon1.
+ apply le_max_r.
+ constructor.
+ apply H.
+ assumption.
+ generalize dependent HA.
+ apply Kont_invar.
+ intros w''.
+ apply sforces'_mon2.
+ apply le_max_l.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_All : forall w A, sforces w (All A) ->
+ forall w', w <= w' -> forall d, locli d -> Kont sforces w' (A^^d).
+ Proof.
+ intros w A H.
+ unfold sforces in H.
+ simpl in H.
+ intros w' ww' d Hd.
+
+ change (Kont (sforces' (depth (A^^d))) w' (A^^d)).
+ assert (H' := H w' ww' d Hd).
+ generalize dependent H'.
+ apply Kont_invar.
+ apply sforces'_mon2.
+ rewrite depth_subst.
+ constructor.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_All' : forall w A,
+ (forall w', w <= w' -> forall d, locli d -> Kont sforces w' (A^^d))
+ -> sforces w (All A).
+ Proof.
+ intros w A H.
+ unfold sforces.
+ simpl.
+ intros w' ww' d Hd.
+ rewrite <- depth_subst with A d 0.
+ apply H.
+ assumption.
+ assumption.
+ Defined.
+
+ Lemma sforces_correct_Ex : forall w A, sforces w (Ex A) ->
+ sigT (fun d => (locli d) * Kont sforces w (A^^d)).
+ Proof.
+ intros w A H.
+ unfold sforces in H.
+ simpl in H.
+ destruct H as [d [Hd H]].
+
+ exists d.
+ split; [assumption|].
+ change (Kont (sforces' (depth (A^^d))) w (A^^d)).
+ generalize dependent H.
+ apply Kont_invar.
+ apply sforces'_mon2.
+ rewrite depth_subst.
+ constructor.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_Ex' : forall w A,
+ (sigT (fun d => (locli d) * Kont sforces w (A^^d)))
+ -> sforces w (Ex A).
+ Proof.
+ intros w A H.
+ unfold sforces.
+ simpl.
+ destruct H as [d [Hd H]].
+ exists d.
+ split; [assumption|].
+ rewrite <- depth_subst with A d 0.
+ apply H.
+ Defined.
+ End sforces_correctness.
+
+ Definition refutes (w1:K)(A:typ) :=
+ forall w2:K, w1 <= w2 -> sforces w2 A -> exploding w2.
+
+ Notation "w : ||+ A " := (sforces w A) (at level 70).
+ Notation "w : A ||- " := (refutes w A) (at level 70).
+ Notation "w : ||- A " := (Kont sforces w A) (at level 70).
+
+ Fixpoint refutes_cxt (w:K)(Delta:cxt_ect) : Type :=
+ match Delta with
+ | nil => unit
+ | cons aA Delta' => (refutes w (snd aA)) * (refutes_cxt w Delta')
+ end.
+
+ Fixpoint forces_cxt (w:K)(Gamma:cxt_trm) : Type :=
+ match Gamma with
+ | nil => unit
+ | cons aA Gamma' => (w : ||- (snd aA)) * (forces_cxt w Gamma')
+ end.
+
+ Notation " w : ||-- Gamma" := (forces_cxt w Gamma) (at level 70).
+ Notation " w : Delta ||-- " := (refutes_cxt w Delta) (at level 70).
+
+ Lemma sforces_mon : forall A w, w : ||+ A -> forall w', w <= w' -> w' : ||+ A .
+ Proof.
+ induction A.
+
+ unfold sforces.
+ simpl.
+ intros.
+ eauto using aforces_mon.
+
+ unfold sforces.
+ intros w H1 w' ww'.
+ simpl in *.
+ eauto using wle_trans.
+
+ unfold sforces.
+ intros w H1 w' ww'.
+ simpl in *.
+ case H1; intro H'.
+ left.
+ eauto using wle_trans.
+ eauto using wle_trans.
+
+ unfold sforces.
+ intros w H1 w' ww'.
+ simpl in *.
+ destruct H1; split; eauto using wle_trans.
+
+ unfold sforces.
+ intros w H1 w' ww'.
+ simpl in *.
+ eauto using wle_trans.
+
+ unfold sforces.
+ intros w H1 w' ww'.
+ simpl in *.
+ destruct H1 as [d [Hd H]].
+ exists d.
+ split; [assumption|].
+ intros w1 w'w1 k.
+ apply H.
+ eauto using wle_trans.
+ intros w2 w1w2 H2.
+ auto.
+ Defined.
+
+ Definition ret {w A} : w : ||+ A -> w : ||- A.
+ Proof.
+ intros H.
+ intros w1 w_w1 k.
+ apply k.
+ apply wle_refl.
+ apply sforces_mon with w.
+ assumption.
+ assumption.
+ Defined.
+
+ Lemma refutes_mon : forall A w, w : A ||- -> forall w', w <= w' -> w' : A ||- .
+ Proof.
+ induction A;
+ intros;
+ unfold refutes in *;
+ eauto using wle_trans.
+ Defined.
+
+ Lemma forces_mon : forall A w, w : ||- A -> forall w', w <= w' -> w' : ||- A.
+ Proof.
+ induction A;
+ intros;
+ eauto using wle_trans.
+ Defined.
+
+ Lemma refutes_cxt_mon :
+ forall Delta w, w : Delta ||-- -> forall w', w <= w' -> w' : Delta ||-- .
+ Proof.
+ induction Delta.
+ simpl.
+ auto.
+ simpl.
+ intros.
+ destruct X.
+ split; eauto using wle_trans,refutes_mon.
+ Defined.
+
+ Lemma forces_cxt_mon :
+ forall Gamma w, w : ||-- Gamma -> forall w', w <= w' -> w' : ||-- Gamma.
+ Proof.
+ induction Gamma.
+ simpl.
+ auto.
+ simpl.
+ intros.
+ destruct X.
+ split; eauto using wle_trans,forces_mon.
+ Defined.
+
+ Lemma refutes_cxt_In : forall w alpha A Delta,
+ In (alpha, A) Delta -> w : Delta ||-- -> w : A ||- .
+ Proof.
+ induction Delta.
+ simpl.
+ intros.
+ contradict H.
+ simpl.
+ intros H H0.
+ destruct H.
+ rewrite e in H0.
+ intuition.
+ intuition.
+ Defined.
+
+ Lemma forces_cxt_In : forall w x A Gamma,
+ In (x, A) Gamma -> w : ||-- Gamma -> w : ||- A.
+ Proof.
+ induction Gamma.
+ simpl.
+ intros.
+ contradict H.
+ simpl.
+ intros H H0.
+ destruct H.
+ rewrite e in H0.
+ intuition.
+ intuition.
+ Defined.
+
+ Theorem soundness :
+ (forall c Gamma Delta, c:(Gamma|-Delta) ->
+ forall w, w : ||-- Gamma -> w : Delta ||-- -> exploding w) *
+
+ (forall Gamma t A Delta, Gamma|-(t:A);Delta ->
+ forall w, w : ||-- Gamma -> w : Delta ||-- -> w : ||- A) *
+
+ (forall Gamma e A Delta, Gamma;(e:A)|-Delta ->
+ forall w, w : ||-- Gamma -> w : Delta ||-- -> w : A ||- ).
+ Proof.
+ apply proof_rect'.
+
+ (* Cut *)
+ intros v e Gamma Delta A _ IH1 _ IH2.
+ intros w HGamma HDelta.
+ apply IH1 with w; auto.
+ apply wle_refl.
+ apply IH2; assumption.
+
+ (* AxR *)
+ intros.
+ eauto using forces_cxt_In.
+
+ (* Mu *)
+ intros Gamma alpha c A Delta _ IH1 w H H0.
+ intros w' ww' H1.
+ apply IH1.
+ eauto using forces_cxt_mon,wle_trans.
+ simpl.
+ eauto using refutes_cxt_mon,refutes_mon.
+
+ (* ImpR *)
+ intros.
+ apply ret.
+ apply sforces_correct_Imp'.
+ simpl in *.
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+
+ (* DisjRL *)
+ intros.
+ apply ret.
+ apply sforces_correct_Disj'.
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+
+ (* DisjRR *)
+ intros.
+ apply ret.
+ apply sforces_correct_Disj'.
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+
+ (* ConjR *)
+ intros.
+ apply ret.
+ apply sforces_correct_Conj'.
+ split;
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+
+ (* AllR *)
+ intros.
+ apply ret.
+ apply sforces_correct_All'.
+ admit.
+
+ (* ExR *)
+ intros.
+ apply ret.
+ apply sforces_correct_Ex'.
+ exists d.
+ split; [assumption|].
+ auto.
+
+ (* Ax_L *)
+ intros.
+ eauto using refutes_cxt_In.
+
+ (* MuT *)
+ intros Gamma a c A Delta _ IH1.
+ intros w H H0.
+ intros w' ww' HA.
+ apply IH1.
+ simpl; split.
+ apply ret; assumption.
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+
+ (* ImpL *)
+ intros Gamma v e A B Delta _ IH1 _ IH2 w HGamma HDelta.
+ intros w' ww' H.
+ eapply sforces_correct_Imp in H.
+ apply H.
+ apply wle_refl.
+ apply IH2;
+ eauto using refutes_cxt_mon,forces_cxt_mon.
+ apply wle_refl.
+ eauto using forces_mon,refutes_cxt_mon,forces_cxt_mon.
+
+ (* DisjL *)
+ intros Gamma e1 e2 A1 A2 Delta _ IH1 _ IH2.
+ intros w H H0.
+ intros w' ww' HDisj.
+ apply sforces_correct_Disj in HDisj.
+ case HDisj; intro Hcase.
+ apply Hcase.
+ apply wle_refl.
+ apply IH1; eauto using refutes_cxt_mon,forces_cxt_mon.
+ apply Hcase.
+ apply wle_refl.
+ apply IH2; eauto using refutes_cxt_mon,forces_cxt_mon.
+
+ (* ConjLL *)
+ intros Gamma e A1 A2 Delta _ IH w H H0.
+ intros w' H1 H2.
+ apply sforces_correct_Conj in H2.
+ destruct H2 as [H2 _].
+ apply H2.
+ apply wle_refl.
+ apply IH; eauto using refutes_cxt_mon,forces_cxt_mon.
+
+ (* ConjLR *)
+ intros Gamma e A1 A2 Delta _ IH w H H0.
+ intros w' H1 H2.
+ apply sforces_correct_Conj in H2.
+ destruct H2 as [_ H2].
+ apply H2.
+ apply wle_refl.
+ apply IH; eauto using refutes_cxt_mon,forces_cxt_mon.
+
+ (* AllL *)
+ intros Gamma e A1 A2 d Hd _ IH w H H0.
+ intros w' H1 H2.
+ eapply sforces_correct_All in H2.
+ apply H2.
+ apply wle_refl.
+ apply IH; eauto using refutes_cxt_mon,forces_cxt_mon.
+ apply wle_refl.
+ assumption.
+
+ (* ExL *)
+ intros Gamma e A Delta L Hlocl _ IH w H H0.
+ intros w' H1 H2.
+ admit.
+ Defined.
+End Abstract_Semantics.
+
+(** Various lemmas that have to do with substitutions related to quantifier bound variables; this was the hardest part of the formalisation. *)
+Section subst_lemmas.
+
+ Lemma locli_fvar : forall x, locli (fvar x).
+ Proof.
+ unfold locli.
+ simpl.
+ reflexivity.
+ Defined.
+
+ Lemma func_wrapper1 : forall f l l', func f l = func f l' -> l = l'.
+ Proof.
+ intros.
+ congruence.
+ Defined.
+
+ Lemma func_wrapper2 : forall f l l', l = l' -> func f l = func f l'.
+ Proof.
+ intros.
+ congruence.
+ Defined.
+
+ Lemma Atom_wrapper1 : forall f l l', Atom f l = Atom f l' -> l = l'.
+ Proof.
+ intros.
+ congruence.
+ Defined.
+
+ Lemma Atom_wrapper2 : forall p l l', l = l' -> Atom p l = Atom p l'.
+ Proof.
+ intros.
+ congruence.
+ Defined.
+
+ Lemma cons_eqs : forall T a a' l l', @cons T a l = @cons T a' l' ->
+ (a=a') * (l=l').
+ Proof.
+ intros.
+ inversion H.
+ auto.
+ Defined.
+
+ Lemma eqs_cons : forall T a a' l l',
+ (a=a') * (l=l') -> @cons T a l = @cons T a' l'.
+ Proof.
+ intros.
+ destruct H.
+ rewrite e.
+ rewrite e0.
+ auto.
+ Defined.
+
+ Lemma subst_indiv_zero_twice : forall a d d' n', locli d ->
+ subst_indiv n' (subst_indiv n' a d) d' = subst_indiv n' a d.
+ Proof.
+ fix 1.
+ intro a.
+ destruct a.
+
+ intros.
+ simpl.
+ assert (H' := beq_nat_false n' n).
+ destruct beq_nat.
+ apply H.
+ simpl.
+ assert (H'' := beq_nat_true n' n).
+ destruct beq_nat.
+ assert (False).
+ apply H'.
+ reflexivity.
+ apply H''.
+ reflexivity.
+ intuition.
+ reflexivity.
+
+ simpl.
+ reflexivity.
+
+ simpl.
+ intros.
+ apply func_wrapper2.
+ induction l.
+ reflexivity.
+ apply eqs_cons.
+ split.
+ apply subst_indiv_zero_twice.
+ assumption.
+
+ apply IHl.
+ Defined.
+
+ Lemma locli_locli_subst : forall a m k d c,
+ (forall k c, subst_indiv (S (k+m)) a c = a) ->
+ (forall k c, subst_indiv (k+m) d c = d) ->
+ subst_indiv (k+m) (subst_indiv m a d) c = subst_indiv m a d.
+ Proof.
+ fix 1.
+ destruct a.
+
+ intros.
+ simpl.
+ remember (beq_nat m n).
+ destruct b.
+ apply H0.
+ destruct k.
+ simpl.
+ rewrite <- Heqb.
+ reflexivity.
+ rewrite plus_Sn_m.
+ apply H.
+
+ simpl.
+ reflexivity.
+
+ intros.
+ simpl.
+ simpl in H.
+ apply func_wrapper2.
+ assert (H' := fun k d => func_wrapper1 (H k d)).
+ clear H.
+ induction l.
+ reflexivity.
+ apply eqs_cons.
+ assert (H1 := fun k d => fst (cons_eqs (H' k d))).
+ assert (H2 := fun k d => snd (cons_eqs (H' k d))).
+ clear H'.
+ split.
+ rewrite locli_locli_subst.
+ reflexivity.
+ apply H1.
+ assumption.
+ apply IHl.
+ apply H2.
+ Defined.
+
+ Lemma locl_locli_subst : forall A m,
+ (forall k d, subst (S (k + m)) A d = A) ->
+ forall d,
+ (forall k d0, subst_indiv (k + m) d d0 = d) ->
+ forall k' d',
+ subst (k'+m) (subst m A d) d' = subst m A d.
+ Proof.
+ fix 1.
+ intros.
+ rename H into H'.
+ destruct A.
+
+ simpl in *.
+ apply Atom_wrapper2.
+ assert (H'' := fun k d => Atom_wrapper1 (H' k d)).
+ clear H'.
+ induction l.
+ reflexivity.
+ assert (H1 := fun k d => fst (cons_eqs (H'' k d))).
+ assert (H2 := fun k d => snd (cons_eqs (H'' k d))).
+ clear H''.
+ apply eqs_cons.
+ split.
+ rewrite locli_locli_subst.
+ reflexivity.
+ apply H1.
+ apply H0.
+ apply IHl.
+ apply H2.
+
+ simpl in *.
+ rewrite locl_locli_subst.
+ rewrite locl_locli_subst.
+ reflexivity.
+ unfold locl.
+ simpl.
+ intros.
+ assert (H'' := H' k d0).
+ congruence.
+ assumption.
+ unfold locl.
+ simpl.
+ intros.
+ assert (H'' := H' k d0).
+ congruence.
+ assumption.
+
+ simpl in *.
+ rewrite locl_locli_subst.
+ rewrite locl_locli_subst.
+ reflexivity.
+ unfold locl.
+ simpl.
+ intros.
+ assert (H'' := H' k d0).
+ congruence.
+ assumption.
+ unfold locl.
+ simpl.
+ intros.
+ assert (H'' := H' k d0).
+ congruence.
+ assumption.
+
+ simpl in *.
+ rewrite locl_locli_subst.
+ rewrite locl_locli_subst.
+ reflexivity.
+ unfold locl.
+ simpl.
+ intros.
+ assert (H'' := H' k d0).
+ congruence.
+ assumption.
+ unfold locl.
+ simpl.
+ intros.
+ assert (H'' := H' k d0).
+ congruence.
+ assumption.
+
+ simpl in *.
+ rewrite plus_n_Sm.
+ rewrite locl_locli_subst.
+ reflexivity.
+ intros.
+ assert (H'' := H' k d0).
+ rewrite <- plus_n_Sm.
+ congruence.
+ intros.
+ rewrite <- Arith.Plus.plus_Snm_nSm.
+ apply H0.
+
+ simpl in *.
+ rewrite plus_n_Sm.
+ rewrite locl_locli_subst.
+ reflexivity.
+ intros.
+ assert (H'' := H' k d0).
+ rewrite <- plus_n_Sm.
+ congruence.
+ intros.
+ rewrite <- Arith.Plus.plus_Snm_nSm.
+ apply H0.
+ Defined.
+
+ Lemma locl_locli_subst' : forall A, locl (All A) ->
+ forall d, locli d -> locl (A^^d).
+ Proof.
+ unfold locl,locli.
+ simpl.
+ intros.
+ replace k with (plus k 0).
+ apply locl_locli_subst.
+ intros.
+ assert (H' := H (plus k0 0) d1).
+ congruence.
+ intros.
+ apply H0.
+ apply Arith.Plus.plus_0_r.
+ Defined.
+
+ Lemma locl_locli_subst'' : forall A, locl (Ex A) ->
+ forall d, locli d -> locl (A^^d).
+ Proof.
+ unfold locl,locli.
+ simpl.
+ intros.
+ replace k with (plus k 0).
+ apply locl_locli_subst.
+ intros.
+ assert (H' := H (plus k0 0) d1).
+ congruence.
+ intros.
+ apply H0.
+ apply Arith.Plus.plus_0_r.
+ Defined.
+
+ Fixpoint lsubst (T:Type)(k:nat)(l:list (T*typ))(d:indiv) :=
+ match l with
+ | cons aA aAs => cons (fst aA, (subst k (snd aA) d)) (lsubst k aAs d)
+ | nil => nil
+ end.
+
+ Notation "l \ x" := (lsubst 0 l (fvar x)) (at level 60).
+
+ Fixpoint fvar_swap_indiv (x:var_free)(a:indiv)(y:var_free) {struct a} : indiv :=
+ let fvar_swap_list := fix fvar_swap_list
+ (l:list indiv) {struct l} : list indiv :=
+ match l with
+ | nil => nil
+ | cons a' l' => cons (fvar_swap_indiv x a' y) (fvar_swap_list l')
+ end in
+ match a with
+ | func f ld =>
+ func f (fvar_swap_list ld)
+ | fvar z => if (var_free_dec x z) then (fvar y) else (fvar z)
+ | bvar i => bvar i
+ end.
+
+ Fixpoint fvar_swap_typ (x:var_free)(A:typ)(y:var_free) {struct A} : typ :=
+ let fvar_swap_list := fix fvar_swap_list
+ (l:list indiv) {struct l} : list indiv :=
+ match l with
+ | nil => nil
+ | cons a' l' => cons (fvar_swap_indiv x a' y) (fvar_swap_list l')
+ end in
+ match A with
+ | Ex A1 => Ex (fvar_swap_typ x A1 y)
+ | All A1 => All (fvar_swap_typ x A1 y)
+ | Conj A1 A2 => Conj (fvar_swap_typ x A1 y) (fvar_swap_typ x A2 y)
+ | Disj A1 A2 => Disj (fvar_swap_typ x A1 y) (fvar_swap_typ x A2 y)
+ | Imp A1 A2 => Imp (fvar_swap_typ x A1 y) (fvar_swap_typ x A2 y)
+ | Atom P ld =>
+ Atom P (fvar_swap_list ld)
+ end.
+
+ Fixpoint fvar_swap_cxt (T:Type)(x:var_free)(l:list (T*typ))(y:var_free) :=
+ match l with
+ | cons aA aAs =>
+ cons (fst aA, (fvar_swap_typ x (snd aA) y)) (fvar_swap_cxt x aAs y)
+ | nil => nil
+ end.
+
+ Lemma locli_fvar_swap : forall a m,
+ (forall k d, subst_indiv (plus k m) a d = a) ->
+ forall x y k d,
+ subst_indiv (plus k m) (fvar_swap_indiv x a y) d = fvar_swap_indiv x a y.
+ Proof.
+ fix 1.
+ intro a.
+ destruct a.
+
+ simpl.
+ auto.
+
+ simpl.
+ intros.
+ destruct var_free_dec.
+ Hint Unfold locli_fvar.
+ apply locli_fvar.
+ apply locli_fvar.
+
+ intros.
+ simpl in *.
+ assert (H' := fun k d => func_wrapper1 (H k d)).
+ clear H.
+ apply func_wrapper2.
+ induction l.
+ reflexivity.
+ apply eqs_cons.
+ assert (H1 := fun k d => fst (cons_eqs (H' k d))).
+ assert (H2 := fun k d => snd (cons_eqs (H' k d))).
+ clear H'.
+ split.
+ unfold locli in locli_fvar_swap.
+ rewrite locli_fvar_swap.
+ reflexivity.
+ apply H1.
+ apply IHl.
+ apply H2.
+ Defined.
+
+ Lemma locl_fvar_swap : forall A m,
+ (forall k d, subst (plus k m) A d = A) ->
+ forall x y k d,
+ subst (plus k m) (fvar_swap_typ x A y) d = fvar_swap_typ x A y.
+ Proof.
+ induction A.
+
+ induction l.
+ simpl.
+ auto.
+ intros m H.
+ simpl in H.
+ assert (H' := fun k d => cons_eqs (Atom_wrapper1 (H k d))).
+ clear H.
+ assert (H1 := fun k d => fst (H' k d)).
+ assert (H2 := fun k d => snd (H' k d)).
+ clear H'.
+ intros.
+ simpl.
+ unfold locl.
+ intros.
+ simpl.
+ apply Atom_wrapper2.
+ apply eqs_cons.
+ split.
+ apply locli_fvar_swap.
+ apply H1.
+ simpl in IHl.
+ assert (IHl' := fun m H x y k d => Atom_wrapper1 (IHl m H x y k d)).
+ clear IHl.
+ assert (H2' : (forall (k : nat) (d : indiv),
+ Atom p
+ ((fix subst_list (l : list indiv) : list indiv :=
+ match l with
+ | nil => nil
+ | a' :: l' => subst_indiv (plus k m) a' d :: subst_list l'
+ end) l) = Atom p l)).
+ intros.
+ assert (H2'' := H2 k0 d0).
+ congruence.
+ apply IHl' with m x y k d in H2'.
+ clear IHl'.
+ apply H2'.
+
+ unfold locl.
+ simpl.
+ intros.
+ rewrite IHA1.
+ rewrite IHA2.
+ reflexivity.
+ assert (H' := fun k d => H k d).
+ intros k' d'.
+ assert (H'' := H' k' d').
+ congruence.
+ assert (H' := fun k d => H k d).
+ intros k' d'.
+ assert (H'' := H' k' d').
+ congruence.
+
+ unfold locl.
+ simpl.
+ intros.
+ rewrite IHA1.
+ rewrite IHA2.
+ reflexivity.
+ assert (H' := fun k d => H k d).
+ intros k' d'.
+ assert (H'' := H' k' d').
+ congruence.
+ assert (H' := fun k d => H k d).
+ intros k' d'.
+ assert (H'' := H' k' d').
+ congruence.
+
+ unfold locl.
+ simpl.
+ intros.
+ rewrite IHA1.
+ rewrite IHA2.
+ reflexivity.
+ assert (H' := fun k d => H k d).
+ intros k' d'.
+ assert (H'' := H' k' d').
+ congruence.
+ assert (H' := fun k d => H k d).
+ intros k' d'.
+ assert (H'' := H' k' d').
+ congruence.
+
+ unfold locl.
+ simpl.
+ intros.
+ rewrite plus_n_Sm.
+ rewrite IHA.
+ reflexivity.
+ intros k' d'.
+ rewrite <- plus_n_Sm.
+ assert (H'' := H k' d').
+ congruence.
+
+ unfold locl.
+ simpl.
+ intros.
+ rewrite plus_n_Sm.
+ rewrite IHA.
+ reflexivity.
+ intros k' d'.
+ rewrite <- plus_n_Sm.
+ assert (H'' := H k' d').
+ congruence.
+ Defined.
+
+ Lemma subst_fvar_swap_indiv : forall i k d x y,
+ fvar_swap_indiv x (subst_indiv k i d) y
+ = subst_indiv k (fvar_swap_indiv x i y) (fvar_swap_indiv x d y).
+ Proof.
+ fix 1.
+ intro i.
+ destruct i.
+
+ simpl.
+ intros.
+ destruct beq_nat.
+ reflexivity.
+ reflexivity.
+
+ simpl.
+ intros.
+ destruct var_free_dec.
+ reflexivity.
+ reflexivity.
+
+ intros.
+ simpl.
+ apply func_wrapper2.
+ induction l.
+ reflexivity.
+ apply eqs_cons.
+ split.
+ rewrite subst_fvar_swap_indiv.
+ reflexivity.
+ apply IHl.
+ Defined.
+
+ Lemma subst_fvar_swap : forall A k d x y, fvar_swap_typ x (subst k A d) y
+ = subst k (fvar_swap_typ x A y) (fvar_swap_indiv x d y).
+ Proof.
+ fix 1.
+ intro A.
+ destruct A.
+
+ intros.
+ simpl.
+
+ apply Atom_wrapper2.
+ induction l.
+ reflexivity.
+ apply eqs_cons.
+ split.
+ rewrite subst_fvar_swap_indiv.
+ reflexivity.
+ apply IHl.
+
+ simpl.
+ intros.
+ rewrite subst_fvar_swap.
+ rewrite subst_fvar_swap.
+ reflexivity.
+
+ simpl.
+ intros.
+ rewrite subst_fvar_swap.
+ rewrite subst_fvar_swap.
+ reflexivity.
+
+ simpl.
+ intros.
+ rewrite subst_fvar_swap.
+ rewrite subst_fvar_swap.
+ reflexivity.
+
+ simpl.
+ intros.
+ rewrite subst_fvar_swap.
+ reflexivity.
+
+ simpl.
+ intros.
+ rewrite subst_fvar_swap.
+ reflexivity.
+ Defined.
+
+ Fixpoint fvar_appears_indiv (x:var_free)(a:indiv) {struct a} : bool :=
+ let fvar_appears_list :=
+ fix fvar_appears_list (x:var_free)(l:list indiv) {struct l} : bool :=
+ match l with
+ | nil => false
+ | cons d l' =>
+ orb (fvar_appears_indiv x d) (fvar_appears_list x l')
+ end
+ in match a with
+ | func f ld => fvar_appears_list x ld
+ | fvar z => if (var_free_dec x z) then true else false
+ | bvar i => false
+ end.
+
+ Fixpoint fvar_appears_typ (x:var_free)(A:typ) {struct A} : bool :=
+ let fvar_appears_list :=
+ fix fvar_appears_list (x:var_free)(l:list indiv) {struct l} : bool :=
+ match l with
+ | nil => false
+ | cons d l' =>
+ orb (fvar_appears_indiv x d) (fvar_appears_list x l')
+ end
+ in match A with
+ | Ex A1 => fvar_appears_typ x A1
+ | All A1 => fvar_appears_typ x A1
+ | Conj A1 A2 => orb (fvar_appears_typ x A1) (fvar_appears_typ x A2)
+ | Disj A1 A2 => orb (fvar_appears_typ x A1) (fvar_appears_typ x A2)
+ | Imp A1 A2 => orb (fvar_appears_typ x A1) (fvar_appears_typ x A2)
+ | Atom P ld => fvar_appears_list x ld
+ end.
+
+ Fixpoint fvar_appears_cxt (T:Type)(x:var_free)(l:list (T*typ)) : bool :=
+ match l with
+ | cons aA aAs =>
+ orb (fvar_appears_typ x (snd aA)) (fvar_appears_cxt x aAs)
+ | nil => false
+ end.
+
+ Lemma swap_appears_not_indiv : forall d x,
+ fvar_appears_indiv x d = false -> forall y, fvar_swap_indiv x d y = d.
+ Proof.
+ fix 1.
+ intro d.
+ destruct d.
+
+ simpl.
+ auto.
+
+ simpl.
+ intros.
+ destruct var_free_dec.
+ discriminate.
+ reflexivity.
+
+ intros.
+ simpl in *.
+ apply func_wrapper2.
+ induction l.
+ reflexivity.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ apply eqs_cons.
+ split.
+ apply swap_appears_not_indiv.
+ assumption.
+ apply IHl.
+ assumption.
+ Defined.
+
+ Lemma swap_appears_not_typ : forall A x,
+ fvar_appears_typ x A = false -> forall y, fvar_swap_typ x A y = A.
+ Proof.
+ fix 1.
+ intro A.
+ destruct A.
+
+ intros.
+ simpl in *.
+ apply Atom_wrapper2.
+ induction l.
+ reflexivity.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ apply eqs_cons.
+ split.
+ apply swap_appears_not_indiv.
+ assumption.
+ apply IHl.
+ assumption.
+
+ simpl.
+ intros.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ rewrite swap_appears_not_typ.
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ assumption.
+ assumption.
+
+ simpl.
+ intros.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ rewrite swap_appears_not_typ.
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ assumption.
+ assumption.
+
+ simpl.
+ intros.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ rewrite swap_appears_not_typ.
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ assumption.
+ assumption.
+
+ simpl.
+ intros.
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ assumption.
+
+ simpl.
+ intros.
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ assumption.
+ Defined.
+
+ Lemma swap_appears_not_cxt : forall T, forall GD:list (T*typ), forall x,
+ fvar_appears_cxt x GD = false -> forall y, fvar_swap_cxt x GD y = GD.
+ Proof.
+ induction GD.
+
+ reflexivity.
+
+ simpl.
+ intros.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ rewrite swap_appears_not_typ; auto.
+ rewrite IHGD;auto.
+ destruct a; auto.
+ Defined.
+
+ Fixpoint FV_indiv (d:indiv) {struct d} : list var_free :=
+ let FV_indiv_list :=
+ fix FV_indiv_list (l:list indiv) {struct l} : list var_free :=
+ match l with
+ | nil => nil
+ | cons d' l' => List.app (FV_indiv d') (FV_indiv_list l')
+ end in
+ match d with
+ | func f vl => FV_indiv_list vl
+ | fvar z => (z::nil)
+ | bvar i => nil
+ end.
+
+ Fixpoint FV_typ (A:typ) {struct A} : list var_free :=
+ let FV_indiv_list :=
+ fix FV_indiv_list (l:list indiv) {struct l} : list var_free :=
+ match l with
+ | nil => nil
+ | cons d' l' => List.app (FV_indiv d') (FV_indiv_list l')
+ end in
+ match A with
+ | Ex A1 => FV_typ A1
+ | All A1 => FV_typ A1
+ | Conj A1 A2 => (FV_typ A1) ++ (FV_typ A2)
+ | Disj A1 A2 => (FV_typ A1) ++ (FV_typ A2)
+ | Imp A1 A2 => (FV_typ A1) ++ (FV_typ A2)
+ | Atom P vl => FV_indiv_list vl
+ end.
+
+ Fixpoint FV_cxt (T:Type)(l:list (T*typ)) {struct l} : list var_free :=
+ match l with
+ | cons aA aAs => FV_typ (snd aA) ++ (FV_cxt aAs)
+ | nil => nil
+ end.
+
+ Lemma notin_app : forall T:Type, forall x:T, forall L1 L2,
+ x\notin L1++L2 -> (x\notin L1)*(x\notin L2).
+ Proof.
+ intros.
+ split.
+ eauto using in_or_app.
+ eauto using in_or_app.
+ Defined.
+
+ Lemma notin_FV_not_appears_indiv : forall a x, x \notin FV_indiv a ->
+ fvar_appears_indiv x a = false.
+ Proof.
+ fix 1.
+ destruct a.
+
+ simpl.
+ reflexivity.
+
+ simpl.
+ intros.
+ destruct var_free_dec.
+ symmetry in e.
+ intuition.
+ reflexivity.
+
+ intros.
+ induction l.
+ reflexivity.
+ simpl in *.
+ apply Bool.orb_false_intro.
+ apply notin_app in H.
+ destruct H.
+ apply notin_FV_not_appears_indiv.
+ assumption.
+ apply IHl.
+ apply notin_app in H.
+ destruct H.
+ assumption.
+ Defined.
+
+ Lemma notin_FV_not_appears : forall A x, x \notin FV_typ A ->
+ fvar_appears_typ x A = false.
+ Proof.
+ induction A.
+
+ simpl.
+ induction l.
+ reflexivity.
+ intros.
+ apply Bool.orb_false_intro.
+ apply notin_app in H.
+ destruct H as [H _].
+ apply notin_FV_not_appears_indiv.
+ assumption.
+ apply IHl.
+ apply notin_app in H.
+ destruct H as [_ H].
+ assumption.
+
+ simpl.
+ intros.
+ apply notin_app in H.
+ destruct H as [H1 H2].
+ apply Bool.orb_false_intro; auto.
+
+ simpl.
+ intros.
+ apply notin_app in H.
+ destruct H as [H1 H2].
+ apply Bool.orb_false_intro; auto.
+
+ simpl.
+ intros.
+ apply notin_app in H.
+ destruct H as [H1 H2].
+ apply Bool.orb_false_intro; auto.
+
+ simpl.
+ intros.
+ auto.
+
+ simpl.
+ intros.
+ auto.
+ Defined.
+
+ Lemma fvar_swap_cxt_idem : forall T:Type, forall GD:list (T*typ),
+ forall x y, x\notin FV_cxt GD ->
+ fvar_swap_cxt x GD y = GD.
+ Proof.
+ induction GD.
+
+ reflexivity.
+
+ intros.
+ simpl.
+ rewrite IHGD.
+ rewrite swap_appears_not_typ.
+ destruct a; auto.
+ simpl in H.
+ apply notin_FV_not_appears.
+ apply notin_app in H.
+ apply H.
+ simpl in H.
+ apply notin_app in H.
+ apply H.
+ Defined.
+
+ Lemma subst_fvar_swap_lem : forall A,
+ forall x y, x\notin FV_typ A ->
+ fvar_swap_typ x (A^x) y = A^y.
+ Proof.
+ intros.
+ replace (A^y) with ((fvar_swap_typ x A y)^^(fvar_swap_indiv x (fvar x) y)).
+ apply subst_fvar_swap.
+ simpl.
+ destruct var_free_dec; [idtac|congruence].
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ apply notin_FV_not_appears.
+ assumption.
+ Defined.
+
+ Definition fvar_swap_proof : forall (x y:var_free),
+ (forall c Gamma Delta, proof_cmd c Gamma Delta ->
+ let Gamma' := fvar_swap_cxt x Gamma y in
+ let Delta' := fvar_swap_cxt x Delta y in
+ proof_cmd c Gamma' Delta') *
+ (forall Gamma t A Delta, proof_trm Gamma t A Delta ->
+ let Gamma' := fvar_swap_cxt x Gamma y in
+ let Delta' := fvar_swap_cxt x Delta y in
+ let A' := fvar_swap_typ x A y in
+ proof_trm Gamma' t A' Delta') *
+ (forall Gamma e A Delta, proof_ect Gamma e A Delta ->
+ let Gamma' := fvar_swap_cxt x Gamma y in
+ let Delta' := fvar_swap_cxt x Delta y in
+ let A' := fvar_swap_typ x A y in
+ proof_ect Gamma' e A' Delta').
+ Proof.
+ intros x y.
+ apply proof_rect'.
+
+ intros v e Gamma Delta A _ IHv _ IHe.
+ apply Cut with (fvar_swap_typ x A y).
+ assumption.
+ assumption.
+
+ intros Gamma a A Delta IH.
+ constructor.
+ clear Delta.
+ induction Gamma.
+ inversion IH.
+ simpl in IH.
+ destruct IH.
+ rewrite e.
+ simpl.
+ auto.
+ simpl.
+ auto.
+
+ intros Gamma alpha c A Delta _ IH.
+ eapply Mu.
+ assumption.
+
+ intros; simpl in *; constructor; auto.
+
+ intros; simpl in *; constructor; auto.
+
+ intros; simpl in *; constructor; auto.
+
+ intros; simpl in *; constructor; auto.
+
+ intros; simpl in *.
+ apply AllR with (x::L).
+ replace (All (fvar_swap_typ x A y)) with (fvar_swap_typ x (All A) y);
+ [idtac | reflexivity].
+ unfold locl.
+ intros.
+ replace k with (plus k 0).
+ apply locl_fvar_swap.
+ intros.
+ apply l.
+ apply Arith.Plus.plus_0_r.
+ intros x0 Hx0.
+ assert (x0 \notin L).
+ clear -Hx0.
+ simpl in *.
+ intro H.
+ auto.
+ replace (fvar_swap_typ x A y ^ x0) with
+ ((fvar_swap_typ x A y) ^^ (fvar_swap_indiv x (fvar x0) y)).
+ rewrite <- subst_fvar_swap.
+ auto.
+ rewrite swap_appears_not_indiv.
+ reflexivity.
+ clear -Hx0.
+ simpl.
+ destruct var_free_dec.
+ simpl in Hx0.
+ assert (Hx0' := Hx0 (inl _ e)).
+ inversion Hx0'.
+ reflexivity.
+
+ intros; simpl in *.
+ apply ExR with (fvar_swap_indiv x d y).
+ unfold locli.
+ intros.
+ replace k with (plus k 0).
+ apply locli_fvar_swap.
+ intros.
+ apply l.
+ apply Arith.Plus.plus_0_r.
+ rewrite <- subst_fvar_swap.
+ assumption.
+
+ intros Gamma alpha A Delta IH.
+ constructor.
+ clear Gamma.
+ induction Delta.
+ inversion IH.
+ simpl in IH.
+ destruct IH.
+ rewrite e.
+ simpl.
+ auto.
+ simpl.
+ auto.
+
+ intros Gamma a c A Delta _ IH.
+ eapply MuT.
+ assumption.
+
+ intros; simpl in *; constructor; auto.
+
+ intros; simpl in *; constructor; auto.
+
+ intros; simpl in *; constructor; auto.
+
+ intros; simpl in *; constructor; auto.
+
+ intros; simpl in *.
+ apply AllL with (fvar_swap_indiv x d y).
+ unfold locli.
+ intros.
+ replace k with (plus k 0).
+ apply locli_fvar_swap.
+ intros.
+ apply l.
+ apply Arith.Plus.plus_0_r.
+ rewrite <- subst_fvar_swap.
+ assumption.
+
+ intros; simpl in *.
+ apply ExL with (x::L).
+ replace (Ex (fvar_swap_typ x A y)) with (fvar_swap_typ x (Ex A) y);
+ [idtac | reflexivity].
+ unfold locl.
+ intros.
+ replace k with (plus k 0).
+ apply locl_fvar_swap.
+ intros.
+ apply l.
+ apply Arith.Plus.plus_0_r.
+ intros x0 Hx0.
+ assert (x0 \notin L).
+ clear -Hx0.
+ simpl in *.
+ intro H.
+ auto.
+ replace (fvar_swap_typ x A y ^ x0) with
+ ((fvar_swap_typ x A y) ^^ (fvar_swap_indiv x (fvar x0) y)).
+ rewrite <- subst_fvar_swap.
+ auto.
+ rewrite swap_appears_not_indiv.
+ reflexivity.
+ clear -Hx0.
+ simpl.
+ destruct var_free_dec.
+ simpl in Hx0.
+ assert (Hx0' := Hx0 (inl _ e)).
+ inversion Hx0'.
+ reflexivity.
+ Defined.
+
+ Lemma proof_trm_quant_invar : forall Gamma Delta A,
+ let L := FV_typ A ++ FV_cxt Gamma ++ FV_cxt Delta in
+ (forall x, x\notin L ->
+ sigT (fun t:trm => proof_trm Gamma t (A^x) Delta)) ->
+ sigT (fun s:trm => forall y, y\notin L -> proof_trm Gamma s (A^y) Delta).
+ Proof.
+ intros.
+ destruct (fresh_fvar L) as [x Hx].
+ destruct (H x) as [t p].
+ assumption.
+ exists t.
+ clear H.
+ unfold L in Hx.
+ apply notin_app in Hx.
+ destruct Hx as [Hx1 Hx2].
+ apply notin_app in Hx2.
+ destruct Hx2 as [Hx2 Hx3].
+ intros y Hy.
+ unfold L in Hy.
+ apply notin_app in Hy.
+ destruct Hy as [Hy1 Hy2].
+ apply notin_app in Hy2.
+ destruct Hy2 as [Hy2 Hy3].
+ assert (p' := (snd (fst (fvar_swap_proof x y))) _ _ _ _ p).
+ simpl in p'.
+ rewrite fvar_swap_cxt_idem in p'; auto.
+ rewrite fvar_swap_cxt_idem in p'; auto.
+ rewrite subst_fvar_swap_lem in p'; auto.
+ Defined.
+
+ Lemma proof_ect_quant_invar : forall Gamma Delta A,
+ let L := FV_typ A ++ FV_cxt Gamma ++ FV_cxt Delta in
+ (forall x, x\notin L ->
+ sigT (fun e:ect => proof_ect Gamma e (A^x) Delta)) ->
+ sigT (fun f:ect => forall y, y\notin L -> proof_ect Gamma f (A^y) Delta).
+ Proof.
+ intros.
+ destruct (fresh_fvar L) as [x Hx].
+ destruct (H x) as [e p].
+ assumption.
+ exists e.
+ clear H.
+ unfold L in Hx.
+ apply notin_app in Hx.
+ destruct Hx as [Hx1 Hx2].
+ apply notin_app in Hx2.
+ destruct Hx2 as [Hx2 Hx3].
+ intros y Hy.
+ unfold L in Hy.
+ apply notin_app in Hy.
+ destruct Hy as [Hy1 Hy2].
+ apply notin_app in Hy2.
+ destruct Hy2 as [Hy2 Hy3].
+ assert (p' := ((snd (fvar_swap_proof x y))) _ _ _ _ p).
+ simpl in p'.
+ rewrite fvar_swap_cxt_idem in p'; auto.
+ rewrite fvar_swap_cxt_idem in p'; auto.
+ rewrite subst_fvar_swap_lem in p'; auto.
+ Defined.
+End subst_lemmas.
+
+(** The Universal model *)
+Section Concrete_Semantics.
+ Definition Kworld : Set := cxt_trm * cxt_ect.
+
+ Definition Kle (w w':Kworld) : Type :=
+ (incl (fst w) (fst w')) * (incl (snd w) (snd w')).
+
+ Lemma Kle_refl : forall w, Kle w w.
+ Proof.
+ intro w.
+ split;
+ auto using incl_refl.
+ Defined.
+
+ Lemma Kle_trans : forall w w' w'', Kle w w' -> Kle w' w'' -> Kle w w''.
+ Proof.
+ unfold Kle.
+ intros.
+ intuition eauto using incl_tran.
+ Defined.
+
+ Definition Kexploding (w:Kworld) : Set :=
+ {c:cmd & proof_cmd c (fst w) (snd w)}.
+ Hint Unfold Kexploding.
+
+ Definition Normal_trm (A:typ)(Gamma:cxt_trm)(Delta:cxt_ect) :=
+ sigT (fun t:trm => proof_trm Gamma t A Delta).
+ Hint Unfold Normal_trm.
+
+ Definition Normal_ect (A:typ)(Gamma:cxt_trm)(Delta:cxt_ect) :=
+ sigT (fun e:ect => proof_ect Gamma e A Delta).
+ Hint Unfold Normal_ect.
+
+ Definition Kaforces (w:Kworld)(p:predicate)(ld:list indiv) : Set :=
+ Normal_trm (@Atom p ld) (fst w) (snd w).
+ Hint Unfold Kaforces.
+
+ Lemma proof_weaken :
+ (forall c Gamma Delta, c:(Gamma|-Delta) ->
+ forall Gamma' Delta', incl Gamma Gamma' -> incl Delta Delta' ->
+ c:(Gamma'|-Delta')) *
+ (forall Gamma t A Delta, Gamma|-(t:A);Delta ->
+ forall Gamma' Delta', incl Gamma Gamma' -> incl Delta Delta' ->
+ Gamma'|-(t:A);Delta') *
+ (forall Gamma e A Delta, Gamma;(e:A)|-Delta ->
+ forall Gamma' Delta', incl Gamma Gamma' -> incl Delta Delta' ->
+ Gamma';(e:A)|-Delta').
+ Proof.
+ apply proof_rect'.
+
+ (* Cut *)
+ eauto using Cut.
+
+ (* AxR *)
+ eauto using AxR.
+
+ (* Mu *)
+ intros.
+ apply Mu.
+ apply X.
+ assumption.
+ red.
+ intuition.
+ simpl in H |- *.
+ intuition.
+
+ (* ImpR *)
+ intros.
+ apply ImpR.
+ apply X.
+ red.
+ simpl.
+ intuition.
+ assumption.
+
+ (* DisjRL *)
+ intros.
+ apply DisjRL.
+ apply X.
+ assumption.
+ assumption.
+
+ (* DisjRR *)
+ intros.
+ apply DisjRR.
+ apply X.
+ assumption.
+ assumption.
+
+ (* ConjR *)
+ intros.
+ apply ConjR.
+ apply X.
+ assumption.
+ assumption.
+ apply X0.
+ assumption.
+ assumption.
+
+ (* AllR *)
+ intros.
+ apply AllR with L.
+ assumption.
+ intros.
+ apply X.
+ assumption.
+ assumption.
+ assumption.
+
+ (* ExR *)
+ intros.
+ apply ExR with d.
+ assumption.
+ apply X.
+ assumption.
+ assumption.
+
+ (* AxL *)
+ eauto using AxL.
+
+ (* MuT *)
+ intros.
+ apply MuT.
+ apply X.
+ red;simpl;intuition.
+ assumption.
+
+ (* ImpL *)
+ eauto using ImpL.
+
+ (* DisjL *)
+ intros.
+ apply DisjL.
+ apply X; intuition.
+ apply X0; intuition.
+
+ (* ConjLL *)
+ intros.
+ apply ConjLL.
+ apply X; intuition.
+
+ (* ConjLR *)
+ intros.
+ apply ConjLR.
+ apply X; intuition.
+
+ (* AllL *)
+ intros.
+ apply AllL with d.
+ assumption.
+ apply X; auto.
+
+ (* ExL *)
+ intros.
+ apply ExL with L.
+ assumption.
+ auto.
+ Defined.
+
+ Lemma Kaforces_mon : forall w P ld, @Kaforces w P ld ->
+ forall w', Kle w w' -> @Kaforces w' P ld.
+ Proof.
+ intros w P ld H.
+ destruct H as [v p].
+ intros w' Hle.
+ exists v.
+ unfold Kle in Hle.
+ destruct Hle.
+ apply (snd (fst proof_weaken)) with (fst w) (snd w); auto.
+ Defined.
+
+ Definition K : Kripke.
+ (* begin show *)
+ apply Build_Kripke with Kworld Kle Kaforces.
+ exact Kle_refl.
+ exact Kle_trans.
+ exact Kexploding.
+ exact Kaforces_mon.
+ (* end show *)
+ Defined.
+End Concrete_Semantics.
+
+Notation "w <= w'" := (Kle w w').
+
+Notation "w : ||+ A " := (@sforces K w A) (at level 70).
+Notation "w : A ||- " := (@refutes K w A) (at level 70).
+Notation "w : ||- A " := (Kont (@sforces K) w A) (at level 70).
+Notation " w : ||-- Gamma " := (@forces_cxt K w Gamma) (at level 70).
+Notation " w : Delta ||-- " := (@refutes_cxt K w Delta) (at level 70).
+
+Definition cmd_to_trm : forall w A, forall alpha:var_ect,
+ let w1 := (fst w, (alpha,A) :: snd w) in
+ sigT (fun c:cmd => proof_cmd c (fst w1) (snd w1)) ->
+ sigT (fun t:trm => proof_trm (fst w) t A (snd w)).
+Proof.
+ intros w A alpha w1 H.
+ destruct H as [c Hc].
+ exists (mu alpha c).
+ constructor.
+ exact Hc.
+Defined.
+
+Definition cmd_to_ect : forall w A, forall x:var_trm,
+ let w1 := ((x,A) :: fst w, snd w) in
+ sigT (fun c:cmd => proof_cmd c (fst w1) (snd w1)) ->
+ sigT (fun e:ect => proof_ect (fst w) e A (snd w)).
+Proof.
+ intros w A x w1 H.
+ destruct H as [c Hc].
+ exists (mut x c).
+ constructor.
+ exact Hc.
+Defined.
+
+Definition to_cmd : forall w A,
+ sigT (fun t:trm => proof_trm (fst w) t A (snd w)) ->
+ sigT (fun e:ect => proof_ect (fst w) e A (snd w)) ->
+ sigT (fun c:cmd => proof_cmd c (fst w) (snd w)).
+Proof.
+ intros w A H1 H2.
+ destruct H1 as [t Ht].
+ destruct H2 as [e He].
+ exists (cm t e).
+ apply Cut with A.
+ assumption.
+ assumption.
+Defined.
+
+Lemma Kle_eta_r : forall w a, Kle w (fst w, a :: snd w).
+Proof.
+ intros.
+ split.
+ destruct w.
+ simpl.
+ auto.
+ simpl.
+ auto.
+Defined.
+
+Lemma Kle_eta_l : forall w a, Kle w (a :: fst w, snd w).
+Proof.
+ intros.
+ split.
+ destruct w.
+ simpl.
+ auto.
+ simpl.
+ auto.
+Defined.
+
+(** The proof of completeness is via a reify-reflect pair, by induction on [n] i.e. by induction on the complexity of the formula [A]. [tfresh] and [efresh] are a fresh variable counters that are increased at every recursive call. *)
+Theorem Kcompleteness :
+ forall (n:nat)(A:typ), le (depth A) n -> locl A ->
+ forall (w:K)(tfresh:var_trm)(efresh:var_ect),
+ (w : ||- A -> Normal_trm A (fst w) (snd w)) *
+ (w : A ||- -> Normal_ect A (fst w) (snd w)).
+Proof.
+ induction n.
+
+ intros A Hn.
+ destruct A;inversion Hn.
+ intros A HSn Hlocl.
+ intros.
+ destruct A.
+
+ (* Atomic formulas *)
+ split.
+
+ intro H.
+ apply cmd_to_trm with efresh.
+ apply H.
+ simpl.
+ apply Kle_eta_r.
+ simpl.
+ intros w1 w1w2.
+ unfold sforces.
+ simpl.
+ intro H1.
+ apply to_cmd with (Atom p l).
+ exact H1.
+ exists (evar efresh).
+ constructor.
+ apply w1w2.
+ simpl.
+ left; reflexivity.
+ intro H.
+ apply cmd_to_ect with tfresh.
+ apply H.
+ apply Kle_eta_l.
+ unfold sforces.
+ simpl.
+ exists (tvar tfresh).
+ constructor.
+ simpl.
+ left; reflexivity.
+
+ (* Implication *)
+ assert (Hdepth2 : le (depth A2) n).
+ clear -HSn.
+ simpl in HSn.
+ apply le_S_n in HSn.
+ eapply le_trans.
+ apply le_max_r.
+ apply HSn.
+ assert (Hdepth1 : le (depth A1) n).
+ clear -HSn.
+ simpl in HSn.
+ apply le_S_n in HSn.
+ eapply le_trans.
+ apply le_max_l.
+ apply HSn.
+ assert (Hlocl1 : locl A1).
+ clear -Hlocl.
+ unfold locl in *.
+ simpl in *.
+ intros.
+ assert (Hlocl' := Hlocl k d).
+ congruence.
+ assert (Hlocl2 : locl A2).
+ clear -Hlocl.
+ unfold locl in *.
+ simpl in *.
+ intros.
+ assert (Hlocl' := Hlocl k d).
+ congruence.
+ split.
+
+ intro H.
+ set (w1 := (((tfresh,A1)::fst w),(snd w))).
+ assert (Normal_trm A2 (fst w1) (snd w1)).
+ apply cmd_to_trm with efresh.
+ apply H.
+ unfold w1.
+ simpl.
+ split.
+ simpl.
+ auto.
+ simpl.
+ auto.
+ intros w3 w2w3 HImp.
+ assert (HImp' := sforces_correct_Imp HImp).
+ clear HImp.
+ assert (HA2 : Normal_trm A2 (fst w3) (snd w3)).
+ apply (IHn _ Hdepth2 Hlocl2 _ (S tfresh) (S efresh)).
+ apply HImp'.
+ apply wle_refl.
+ intros w4 w3w4 HA1.
+ change (w4:A1||-) in HA1.
+ apply (IHn _ Hdepth1 Hlocl1 _ (S tfresh) (S efresh)) in HA1.
+ destruct HA1 as [t1 Ht1].
+ eexists.
+ apply Cut with A1.
+ apply AxR.
+ apply w3w4.
+ apply w2w3.
+ unfold w1.
+ left;reflexivity.
+ apply Ht1.
+ destruct HA2 as [t2 Ht2].
+ exists (cm (t2) (evar efresh)).
+ apply Cut with A2.
+ assumption.
+ constructor.
+ apply w2w3.
+ left;reflexivity.
+ destruct H0 as [t2' Ht2'].
+ eexists.
+ apply ImpR.
+ unfold w1 in Ht2'.
+ simpl in Ht2'.
+ apply Ht2'.
+
+ intro H.
+ apply cmd_to_ect with tfresh.
+ apply H.
+ apply Kle_eta_l.
+ apply sforces_correct_Imp'.
+ intros w2 w1w2 HA1.
+ apply (IHn _ Hdepth1 Hlocl1 _ (S tfresh) (S efresh)) in HA1.
+ intros w3 w2w3 HA2.
+ change (w3:A2||-) in HA2.
+ apply (IHn _ Hdepth2 Hlocl2 _ (S tfresh) (S efresh)) in HA2.
+ destruct HA1 as [t1 Ht1].
+ destruct HA2 as [e2 He2].
+ eexists.
+ apply Cut with (Imp A1 A2).
+ apply AxR.
+ apply w2w3.
+ apply w1w2.
+ left;reflexivity.
+ apply ImpL.
+ apply (snd (fst proof_weaken)) with (fst w2) (snd w2).
+ apply Ht1.
+ destruct w2w3; auto.
+ destruct w2w3; auto.
+ apply He2.
+
+ (* Disjunction *)
+ (* the following 4 assertions are a copy-past of above; to define an Ltac? *)
+ assert (Hdepth2 : le (depth A2) n).
+ clear -HSn.
+ simpl in HSn.
+ apply le_S_n in HSn.
+ eapply le_trans.
+ apply le_max_r.
+ apply HSn.
+ assert (Hdepth1 : le (depth A1) n).
+ clear -HSn.
+ simpl in HSn.
+ apply le_S_n in HSn.
+ eapply le_trans.
+ apply le_max_l.
+ apply HSn.
+ assert (Hlocl1 : locl A1).
+ clear -Hlocl.
+ unfold locl in *.
+ simpl in *.
+ intros.
+ assert (Hlocl' := Hlocl k d).
+ congruence.
+ assert (Hlocl2 : locl A2).
+ clear -Hlocl.
+ unfold locl in *.
+ simpl in *.
+ intros.
+ assert (Hlocl' := Hlocl k d).
+ congruence.
+ split.
+
+ intro H.
+ apply cmd_to_trm with efresh.
+ apply H.
+ apply Kle_eta_r.
+ intros w2 w1w2 H2.
+ apply sforces_correct_Disj in H2.
+ case H2.
+ intro case1.
+ apply (IHn _ Hdepth1 Hlocl1 _ (S tfresh) (S efresh)) in case1.
+ simpl.
+ apply to_cmd with (Disj A1 A2).
+ destruct case1 as [t1 Ht1].
+ exists (injl t1).
+ constructor.
+ exact Ht1.
+ exists (evar efresh).
+ constructor.
+ apply w1w2.
+ simpl.
+ left;reflexivity.
+ intro case2.
+ apply (IHn _ Hdepth2 Hlocl2 _ (S tfresh) (S efresh)) in case2.
+ simpl.
+ apply to_cmd with (Disj A1 A2).
+ destruct case2 as [t2 Ht2].
+ exists (injr t2).
+ constructor.
+ exact Ht2.
+ exists (evar efresh).
+ constructor.
+ apply w1w2.
+ simpl.
+ left;reflexivity.
+
+ intro H.
+ set (w1 := ((tfresh, A1)::fst w, snd w)).
+ assert (ww1 : wle w w1).
+ apply Kle_eta_l.
+ assert (branch1 : Normal_ect A1 (fst w1) (snd w1)).
+ apply (IHn _ Hdepth1 Hlocl1 _ (S tfresh) (S efresh)).
+ intros w' w1w' H'.
+ apply H.
+ eauto using wle_trans.
+ apply sforces_correct_Disj'.
+ left.
+ apply ret.
+ assumption.
+ set (w2 := ((tfresh, A2)::fst w, snd w)).
+ assert (ww2 : wle w w2).
+ apply Kle_eta_l.
+ assert (branch2 : Normal_ect A2 (fst w2) (snd w2)).
+ apply (IHn _ Hdepth2 Hlocl2 _ (S tfresh) (S efresh)).
+ intros w' w2w' H'.
+ apply H.
+ eauto using wle_trans.
+ apply sforces_correct_Disj'.
+ right.
+ apply ret.
+ assumption.
+ destruct branch1 as [e1 He1].
+ destruct branch2 as [e2 He2].
+ exists (disj (mut tfresh (cm (tvar tfresh) e1))
+ (mut tfresh (cm (tvar tfresh) e2))).
+ constructor.
+ constructor.
+ apply Cut with A1.
+ constructor.
+ left;reflexivity.
+ assumption.
+ constructor.
+ apply Cut with A2.
+ constructor.
+ left;reflexivity.
+ assumption.
+
+ (* Conjunction *)
+ (* the following 4 assertions are a copy-past of above; to define an Ltac? *)
+ assert (Hdepth2 : le (depth A2) n).
+ clear -HSn.
+ simpl in HSn.
+ apply le_S_n in HSn.
+ eapply le_trans.
+ apply le_max_r.
+ apply HSn.
+ assert (Hdepth1 : le (depth A1) n).
+ clear -HSn.
+ simpl in HSn.
+ apply le_S_n in HSn.
+ eapply le_trans.
+ apply le_max_l.
+ apply HSn.
+ assert (Hlocl1 : locl A1).
+ clear -Hlocl.
+ unfold locl in *.
+ simpl in *.
+ intros.
+ assert (Hlocl' := Hlocl k d).
+ congruence.
+ assert (Hlocl2 : locl A2).
+ clear -Hlocl.
+ unfold locl in *.
+ simpl in *.
+ intros.
+ assert (Hlocl' := Hlocl k d).
+ congruence.
+ split.
+
+ intro H.
+ assert (conj1 : Normal_trm A1 (fst w) (snd w)).
+ apply cmd_to_trm with efresh.
+ apply H.
+ apply Kle_eta_r.
+ intros w2 w1w2 H2.
+ apply sforces_correct_Conj in H2.
+ destruct H2 as [Ht1 _].
+ apply (IHn _ Hdepth1 Hlocl1 _ (S tfresh) (S efresh)) in Ht1.
+ destruct Ht1 as [t1 Ht1].
+ exists (cm t1 (evar efresh)).
+ apply Cut with A1.
+ assumption.
+ constructor.
+ apply w1w2.
+ left;reflexivity.
+ assert (conj2 : Normal_trm A2 (fst w) (snd w)).
+ apply cmd_to_trm with efresh.
+ apply H.
+ apply Kle_eta_r.
+ intros w2 w1w2 H2.
+ apply sforces_correct_Conj in H2.
+ destruct H2 as [_ Ht2].
+ apply (IHn _ Hdepth2 Hlocl2 _ (S tfresh) (S efresh)) in Ht2.
+ destruct Ht2 as [t2 Ht2].
+ exists (cm t2 (evar efresh)).
+ apply Cut with A2.
+ assumption.
+ constructor.
+ apply w1w2.
+ left;reflexivity.
+ destruct conj1 as [t1' Ht1'].
+ destruct conj2 as [t2' Ht2'].
+ exists (paire t1' t2').
+ constructor.
+ assumption.
+ assumption.
+
+ intro H.
+ apply cmd_to_ect with (tfresh).
+ apply H.
+ apply Kle_eta_l.
+ apply sforces_correct_Conj'.
+ clear H.
+ split.
+ intros w2 w1w2 k.
+ change (w2:A1||-) in k.
+ apply (IHn _ Hdepth1 Hlocl1 _ (S tfresh) (S efresh)) in k.
+ destruct k as [e1 He1].
+ exists (cm (tvar tfresh) (projl e1)).
+ apply Cut with (Conj A1 A2).
+ constructor.
+ apply w1w2.
+ left;reflexivity.
+ constructor.
+ assumption.
+ intros w2 w1w2 k.
+ change (w2:A2||-) in k.
+ apply (IHn _ Hdepth2 Hlocl2 _ (S tfresh) (S efresh)) in k.
+ destruct k as [e2 He2].
+ exists (cm (tvar tfresh) (projr e2)).
+ apply Cut with (Conj A1 A2).
+ constructor.
+ apply w1w2.
+ left;reflexivity.
+ constructor.
+ assumption.
+
+ (* Universal quantifier *)
+ split.
+
+ intro H.
+ set (L := FV_typ A ++ FV_cxt (fst w) ++ FV_cxt (snd w)).
+ assert (H0:forall y:var_free, y \notin L -> Normal_trm (A^y) (fst w) (snd w)).
+ intros.
+ apply cmd_to_trm with efresh.
+ apply H.
+ apply Kle_eta_r.
+ intros w2 w1w2 HAllA.
+ assert (HAllA' := sforces_correct_All HAllA).
+ clear HAllA.
+ assert (Hdepth' : le (depth (A^y)) n).
+ clear -HSn.
+ simpl in HSn.
+ rewrite depth_subst.
+ apply le_S_n in HSn.
+ assumption.
+ assert (Hlocl' : locl (A^y)).
+ apply locl_locli_subst'.
+ assumption.
+ apply locli_fvar.
+ assert (HAd : w2 : ||- A^y).
+ apply HAllA'.
+ clear HAllA'.
+ apply wle_refl.
+ apply locli_fvar.
+ apply (IHn _ Hdepth' Hlocl' _ (S tfresh) (S efresh)) in HAd.
+ destruct HAd as [td Htd].
+ exists (cm td (evar efresh)).
+ apply Cut with (A^y).
+ assumption.
+ apply AxL.
+ apply w1w2.
+ left;reflexivity.
+ apply proof_trm_quant_invar in H0.
+ destruct H0 as [s Hs].
+ exists (tall s).
+ apply AllR with L.
+ assumption.
+ intros.
+ apply Hs.
+ assumption.
+
+ intro H.
+ apply cmd_to_ect with tfresh.
+ apply H.
+ apply Kle_eta_l.
+ apply sforces_correct_All'.
+ intros w2 w1w2 d Hd.
+ intros w3 w2w3 HAd.
+ change (w3:A^^d||-) in HAd.
+ assert (Hdepth' : le (depth (A^^d)) n).
+ clear -HSn.
+ simpl in HSn.
+ rewrite depth_subst.
+ apply le_S_n in HSn.
+ assumption.
+ assert (Hlocl' : locl (A^^d)).
+ apply locl_locli_subst'.
+ assumption.
+ assumption.
+ apply (IHn _ Hdepth' Hlocl' _ (S tfresh) (S efresh)) in HAd.
+ destruct HAd as [ed Hed].
+ exists (cm (tvar tfresh) (eall ed)).
+ apply Cut with (All A).
+ constructor.
+ apply w2w3.
+ apply w1w2.
+ left;reflexivity.
+ apply AllL with d.
+ assumption.
+ assumption.
+
+
+ (* Existential *)
+ split.
+
+ intro H.
+ apply cmd_to_trm with efresh.
+ apply H.
+ apply Kle_eta_r.
+ intros w2 w1w2 H2.
+ apply sforces_correct_Ex in H2.
+ destruct H2 as [d [Hd H2]].
+ assert (Hdepth' : le (depth (A^^d)) n).
+ clear -HSn.
+ simpl in HSn.
+ rewrite depth_subst.
+ apply le_S_n in HSn.
+ assumption.
+ assert (Hlocl' : locl (A^^d)).
+ apply locl_locli_subst''.
+ assumption.
+ assumption.
+ apply (IHn _ Hdepth' Hlocl' _ (S tfresh) (S efresh)) in H2.
+ destruct H2 as [t Ht].
+ simpl.
+ apply to_cmd with (Ex A).
+ exists (tex t).
+ apply ExR with d.
+ assumption.
+ exact Ht.
+ exists (evar efresh).
+ constructor.
+ apply w1w2.
+ simpl.
+ left;reflexivity.
+
+ intro H.
+
+ set (w1 := fun d => ((tfresh, A^^d)::fst w, snd w)).
+ assert (ww1 : forall d, wle w (w1 d)).
+ intro d.
+ apply Kle_eta_l.
+ assert (branch1 : forall d, locli d -> Normal_ect (A^^d) (fst (w1 d)) (snd (w1 d))).
+ intros d Hd.
+ assert (Hdepth' : le (depth (A^^d)) n).
+ clear -HSn.
+ simpl in HSn.
+ rewrite depth_subst.
+ apply le_S_n in HSn.
+ assumption.
+ assert (Hlocl' : locl (A^^d)).
+ apply locl_locli_subst''.
+ assumption.
+ assumption.
+ apply (IHn _ Hdepth' Hlocl' _ (S tfresh) (S efresh)).
+ intros w' w1w' H'.
+ apply H.
+ eauto using wle_trans.
+ apply sforces_correct_Ex'.
+ exists d.
+ split ; [exact Hd|].
+ apply ret.
+ assumption.
+ set (L := FV_typ A ++ FV_cxt (fst w) ++ FV_cxt (snd w)).
+ assert (H0:forall y:var_free, y \notin L -> Normal_ect (A^y) (fst w) (snd w)).
+ intros.
+ destruct (branch1 (fvar y) (locli_fvar _)) as [e He].
+ apply cmd_to_ect with tfresh.
+ simpl.
+ exists (cm (tvar tfresh) e).
+ apply Cut with (A^y).
+ constructor.
+ left;reflexivity.
+ simpl in He.
+ assumption.
+ apply proof_ect_quant_invar in H0.
+ destruct H0 as [f Hf].
+ exists (eex f).
+ apply ExL with L.
+ assumption.
+ intros.
+ apply Hf.
+ assumption.
+Defined.
+
+(** Next: define NbE by composing soundness and completeness. The lemmas [forces_cxt_id] are [refutes_cxt_id] the places completeness is applied at. The variables [tf] and [ef] are the starting fresh variable counters. *)
+
+Fixpoint locl_cxt_trm (Gamma:cxt_trm) : Prop :=
+ match Gamma with
+ | nil => True
+ | cons a Gamma' => locl (snd a) /\ locl_cxt_trm Gamma'
+ end.
+
+Fixpoint locl_cxt_ect (Delta:cxt_ect) : Prop :=
+ match Delta with
+ | nil => True
+ | cons a Delta' => locl (snd a) /\ locl_cxt_ect Delta'
+ end.
+
+Parameter tf : var_trm.
+Parameter ef : var_ect.
+
+Lemma forces_cxt_id : forall Gamma Delta, locl_cxt_trm Gamma -> locl_cxt_ect Delta -> @forces_cxt K (Gamma,Delta) Gamma.
+Proof.
+ induction Gamma.
+ simpl.
+ constructor.
+ simpl.
+ intros Delta [Hlocl HloclGamma] HloclDelta.
+ split.
+ set (w0 := (a :: Gamma, Delta)).
+ intros w1 w0w1 Ha.
+ simpl in w0w1.
+ change (w0 <= w1) in w0w1.
+ change (w1:snd a||-) in Ha.
+ simpl.
+ apply (snd (@Kcompleteness (depth (snd a)) (snd a) (le_n _) Hlocl w1 tf ef)) in Ha.
+ destruct Ha as [e pe].
+ exists (cm (tvar (fst a)) e).
+ apply Cut with (snd a).
+ apply AxR.
+ apply w0w1.
+ simpl.
+ left.
+ destruct a;auto.
+ assumption.
+ apply forces_cxt_mon with (Gamma, Delta).
+ auto.
+ simpl.
+ red.
+ simpl.
+ unfold incl.
+ intuition.
+Defined.
+
+Lemma refutes_cxt_id : forall Gamma Delta, locl_cxt_trm Gamma -> locl_cxt_ect Delta -> @refutes_cxt K (Gamma,Delta) Delta.
+Proof.
+ induction Delta.
+ simpl.
+ constructor.
+ simpl.
+ intros HloclGamma [Hlocl HloclDelta].
+ split.
+ set (w0 := (Gamma, a::Delta)).
+ intros w1 w0w1 Ha.
+ simpl in w0w1.
+ change (w0 <= w1) in w0w1.
+ simpl.
+ apply ret in Ha.
+ apply (fst (@Kcompleteness (depth (snd a)) (snd a) (le_n _) Hlocl w1 tf ef)) in Ha.
+ destruct Ha as [t pt].
+ exists (cm t (evar (fst a))).
+ apply Cut with (snd a).
+ assumption.
+ apply AxL.
+ apply w0w1.
+ simpl.
+ left.
+ destruct a;auto.
+ apply refutes_cxt_mon with (Gamma, Delta).
+ auto.
+ simpl.
+ red.
+ simpl.
+ unfold incl.
+ intuition.
+Defined.
+
+Definition NbE : forall c Gamma Delta,
+ locl_cxt_trm Gamma -> locl_cxt_ect Delta -> proof_cmd c Gamma Delta ->
+ sigT (fun (c':cmd) => proof_cmd c' Gamma Delta).
+Proof.
+ intros c Gamma Delta HloclGamma HloclDelta H.
+ destruct c as [t e].
+ inversion H.
+ clear -H2 H5 HloclGamma HloclDelta.
+ assert (HGamma : @forces_cxt K (Gamma,Delta) Gamma).
+ apply forces_cxt_id; assumption.
+ assert (HDelta : @refutes_cxt K (Gamma,Delta) Delta).
+ apply refutes_cxt_id; assumption.
+ assert (H2' := (snd (fst (soundness K))) Gamma t A Delta H2 (Gamma,Delta) HGamma HDelta).
+ assert (H5' := (snd (soundness K)) Gamma e A Delta H5 (Gamma,Delta) HGamma HDelta).
+ apply H2' in H5'.
+ exact H5'.
+ apply Kle_refl.
+Defined.
+
+
+(** And now some computational tests of normalisation. *)
+
+Parameter P : predicate.
+Parameter li : list indiv.
+Let ATOM := Atom P li.
+Parameter ATOMlocl : locl ATOM.
+
+Parameters x y z : var_trm.
+Parameters alpha beta gamma delta : var_ect.
+
+(** Example 0: a simplest possible beta-reduction (for implication) *)
+Definition cmd0 : cmd := cm (lam x (tvar x)) (app (tvar y) (evar beta)).
+Definition Gamma0 := ((y,ATOM)::nil).
+Definition Delta0 := ((beta,ATOM)::nil).
+Definition proof0 : proof_cmd cmd0 Gamma0 Delta0.
+Proof.
+ apply Cut with (Imp ATOM ATOM).
+ apply ImpR.
+ apply AxR.
+ left; reflexivity.
+ apply ImpL.
+ apply AxR.
+ left; reflexivity.
+ apply AxL.
+ left; reflexivity.
+Defined.
+Definition locl00 : locl_cxt_trm Gamma0.
+Proof.
+ unfold Gamma0.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition locl01 : locl_cxt_ect Delta0.
+Proof.
+ unfold Delta0.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition test0 := projT1 (NbE locl00 locl01 proof0 ).
+(* Time Eval vm_compute in test0. *)
+
+(** Example 1: a seemingly more complex beta-reduction (for implication),
+ which reduces to the same thing as example 0 *)
+Definition cmd1 : cmd := cm (lam x (tvar y)) (app (tvar x) (evar beta)).
+Definition Gamma1 := ((y,ATOM)::(x,ATOM)::nil).
+Definition Delta1 := ((beta,ATOM)::nil).
+Definition proof1 : proof_cmd cmd1 Gamma1 Delta1.
+Proof.
+ apply Cut with (Imp ATOM ATOM).
+ apply ImpR.
+ apply AxR.
+ right.
+ left; reflexivity.
+ apply ImpL.
+ apply AxR.
+ right.
+ left; reflexivity.
+ apply AxL.
+ left; reflexivity.
+Defined.
+Definition locl10 : locl_cxt_trm Gamma1.
+Proof.
+ unfold Gamma1.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition locl11 : locl_cxt_ect Delta1.
+Proof.
+ unfold Delta1.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition test1 := projT1 (NbE locl10 locl11 proof1).
+(* Time Eval vm_compute in test1. *)
+
+(** Example 2: The critical pair indicates we are indeed in
+ call-by-value *)
+Definition c' : cmd := cm (tvar x) (evar gamma).
+Definition c'' : cmd := cm (tvar z) (evar delta).
+Definition cmd2 := cm (mu beta c') (mut y c'').
+Definition Gamma2 := ((x,ATOM)::(z,ATOM)::nil).
+Definition Delta2 := ((gamma,ATOM)::(delta,ATOM)::nil).
+Definition Gamma2' := ((z,ATOM)::(x,ATOM)::nil).
+Definition Delta2' := ((delta,ATOM)::(gamma,ATOM)::nil).
+Definition proof2 : proof_cmd cmd2 Gamma2 Delta2.
+Proof.
+ apply Cut with ATOM.
+ apply Mu.
+ apply Cut with ATOM.
+ apply AxR.
+ left;reflexivity.
+ apply AxL.
+ right;left;reflexivity.
+ apply MuT.
+ apply Cut with ATOM.
+ apply AxR.
+ right;right;left;reflexivity.
+ apply AxL.
+ right;left;reflexivity.
+Defined.
+Definition proof2' : proof_cmd cmd2 Gamma2' Delta2'.
+Proof.
+ apply Cut with ATOM.
+ apply Mu.
+ apply Cut with ATOM.
+ apply AxR.
+ right;left;reflexivity.
+ apply AxL.
+ right;right;left;reflexivity.
+ apply MuT.
+ apply Cut with ATOM.
+ apply AxR.
+ right;left;reflexivity.
+ apply AxL.
+ left;reflexivity.
+Defined.
+Definition locl20 : locl_cxt_trm Gamma2.
+Proof.
+ unfold Gamma2.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition locl21 : locl_cxt_ect Delta2.
+Proof.
+ unfold Delta2.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition locl2'0 : locl_cxt_trm Gamma2'.
+Proof.
+ unfold Gamma2'.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition locl2'1 : locl_cxt_ect Delta2'.
+Proof.
+ unfold Delta2'.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition test2 := projT1 (NbE locl20 locl21 proof2).
+Definition test2' := projT1 (NbE locl2'0 locl2'1 proof2').
+(* Time Eval vm_compute in test2. *)
+(* Time Eval vm_compute in test2'. *)
+
+(** Example 4': verifying reduction for mu *)
+Definition cmd4' := cm (mu gamma (cm (tvar y) (evar gamma))) (evar beta).
+Definition proof4' : proof_cmd cmd4' ((y,ATOM)::nil) ((beta,ATOM)::nil).
+Proof.
+ apply Cut with (ATOM).
+ apply Mu.
+ apply Cut with (ATOM).
+ apply AxR.
+ left;reflexivity.
+ apply AxL.
+ left;reflexivity.
+ apply AxL.
+ left;reflexivity.
+Defined.
+Definition locl40 : locl_cxt_trm ((y,ATOM)::nil).
+Proof.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition locl41 : locl_cxt_ect ((beta,ATOM)::nil).
+Proof.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition test4' := projT1 (NbE locl40 locl41 proof4').
+(* Time Eval vm_compute in test4'. *)
+
+(** Example 5: verifying reduction for conjunction 1 *)
+Definition cmd5 := cm (paire (tvar x) (tvar z)) (projl (evar beta)).
+Definition proof5 : proof_cmd cmd5 ((x,ATOM)::(z,ATOM)::nil) ((beta,ATOM)::nil).
+Proof.
+ apply Cut with (Conj (ATOM) (ATOM)).
+ apply ConjR.
+ apply AxR.
+ left;reflexivity.
+ apply AxR.
+ right;left;reflexivity.
+ apply ConjLL.
+ apply AxL.
+ left;reflexivity.
+Defined.
+Definition locl50 : locl_cxt_trm ((x,ATOM)::(z,ATOM)::nil).
+Proof.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition locl51 : locl_cxt_ect ((beta,ATOM)::nil).
+Proof.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition test5 := projT1 (NbE locl50 locl51 proof5).
+(* Time Eval vm_compute in test5. *)
+
+(** Example 6: verifying reduction for conjunction 2 *)
+Definition cmd6 := cm (paire (tvar x) (tvar z)) (projr (evar beta)).
+Definition proof6 : proof_cmd cmd6 ((x,ATOM)::(z,ATOM)::nil) ((beta,ATOM)::nil).
+Proof.
+ apply Cut with (Conj (ATOM) (ATOM)).
+ apply ConjR.
+ apply AxR.
+ left;reflexivity.
+ apply AxR.
+ right;left;reflexivity.
+ apply ConjLR.
+ apply AxL.
+ left;reflexivity.
+Defined.
+Definition test6 := projT1 (NbE locl50 locl51 proof6).
+(* Time Eval vm_compute in test6. *)
+
+(** Example 7: verifying reduction for disjunction 1 *)
+Definition cmd7 := cm (injl (tvar z)) (disj (evar gamma) (evar delta)).
+Definition proof7 : proof_cmd cmd7 ((z,ATOM)::nil) ((gamma,ATOM)::(delta,ATOM)::nil).
+Proof.
+ apply Cut with (Disj (ATOM) (ATOM)).
+ apply DisjRL.
+ apply AxR.
+ left;reflexivity.
+ apply DisjL.
+ apply AxL.
+ left;reflexivity.
+ apply AxL.
+ right;left;reflexivity.
+Defined.
+Definition locl70 : locl_cxt_trm ((z,ATOM)::nil).
+Proof.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition locl71 : locl_cxt_ect ((gamma,ATOM)::(delta,ATOM)::nil).
+Proof.
+ simpl.
+ auto using ATOMlocl.
+Defined.
+Definition test7 := projT1 (NbE locl70 locl71 proof7).
+(* Time Eval vm_compute in test7. *)
+
+(** Example 8: verifying reduction for disjunction 2 *)
+Definition cmd8 := cm (injr (tvar z)) (disj (evar gamma) (evar delta)).
+Definition proof8 : proof_cmd cmd8 ((z,ATOM)::nil) ((gamma,ATOM)::(delta,ATOM)::nil).
+Proof.
+ apply Cut with (Disj (ATOM) (ATOM)).
+ apply DisjRR.
+ apply AxR.
+ left;reflexivity.
+ apply DisjL.
+ apply AxL.
+ left;reflexivity.
+ apply AxL.
+ right;left;reflexivity.
+Defined.
+Definition test8 := projT1 (NbE locl70 locl71 proof8).
+(* Time Eval vm_compute in test8. *)
+
diff --git a/kripke_completeness/html/index.html b/kripke_completeness/html/index.html
new file mode 100644
index 0000000..d1b6462
--- /dev/null
+++ b/kripke_completeness/html/index.html
@@ -0,0 +1,20 @@
+
+
+Completeness for Kripke-style models formalised in Coq
+
+
+Completeness for Kripke-style models formalised in Coq
+
+The formal proofs are inside the archive kripke_completeness.zip . You can also browse the generated documentation as colored HTML code here .
+
+In order to compile the sources, you need the Coq proof assistant version 8.2.
+
+If you are on a Unix-like system, and the executable coqc is in your PATH variable, you can just run the command make from inside a terminal window. If coqc is not in your path, you can either set it, or open the Coq IDE and run Make from the Compile menu.
+
+If you are on a Windows system, you do not have make , and coqc is not in your PATH, the simplest solution is to open Coq IDE and compile one by one file, making sure to first compile stdlib_Type.v .
+
+
+Back to Danko Ilik's home page
+
+
+
diff --git a/kripke_completeness/int_comp.v b/kripke_completeness/int_comp.v
new file mode 100644
index 0000000..46767dc
--- /dev/null
+++ b/kripke_completeness/int_comp.v
@@ -0,0 +1,2253 @@
+(** Formalised proof of completeness of full Intuitionistic predicate
+ logic with respect to Kripke-style models.
+
+ Danko Ilik, January 2010
+*)
+Require Import stdlib_Type.
+Require Import EqNat.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+(** printing <= $\le$ #≤# *)
+(** printing ||-- $\Vdash$ #⊩# *)
+(** printing ||- $\Vdash$ #⊩# *)
+
+Open Scope type_scope.
+
+Definition var : Set := nat.
+
+(** Proof terms *)
+Inductive term : Set :=
+| Lam : var -> term -> term
+| Inl : term -> term
+| Inr : term -> term
+| Var : var -> term
+| App : term -> term -> term
+| Match : term -> var -> term -> var -> term -> term
+| QLam : term -> term
+| QApp : term -> term
+| QWit : term -> term
+| QMatch : term -> var -> term -> term.
+
+
+(** Abstract (decidable) sets of fresh variables, and predicate and function symbols *)
+Parameter var_free : Set.
+Parameter var_free_dec : forall x y:var_free, {x=y}+{~x=y}.
+Parameters predicate function : Set.
+Notation "x \notin L" := (In x L -> Empty_set) (at level 69).
+Hypothesis fresh_fvar : forall L, sigT (fun x:var_free => x \notin L).
+
+(** Individual terms. [bvar]s are deBruijn indices which represent bound variables, while [fvar]s are names that represent free variables *)
+Inductive indiv : Set :=
+ bvar : nat -> indiv
+| fvar : var_free -> indiv
+| func : function -> list indiv -> indiv.
+
+Inductive formula :=
+ Atom : predicate -> list indiv -> formula
+| Imp : formula -> formula -> formula
+| Disj : formula -> formula -> formula
+| All : formula -> formula
+| Ex : formula -> formula.
+
+(** Substituting at de Bruijn variable k *)
+
+Fixpoint subst_indiv (k:nat)(a:indiv)(d:indiv) {struct a} : indiv :=
+ let subst_list := fix subst_list (l:list indiv) {struct l} : list indiv :=
+ match l with
+ | nil => nil
+ | cons a' l' => cons (subst_indiv k a' d) (subst_list l')
+ end in
+ match a with
+ | func f ld =>
+ func f (subst_list ld)
+ | fvar x => fvar x
+ | bvar i => if beq_nat k i then d else bvar i
+ end.
+
+Fixpoint subst (k:nat)(A:formula)(d:indiv) {struct A} : formula :=
+ let subst_list := fix subst_list (l:list indiv) {struct l} : list indiv :=
+ match l with
+ | nil => nil
+ | cons a' l' => cons (subst_indiv k a' d) (subst_list l')
+ end in
+ match A with
+ | Ex A1 => Ex (subst (S k) A1 d)
+ | All A1 => All (subst (S k) A1 d)
+ | Disj A1 A2 => Disj (subst k A1 d) (subst k A2 d)
+ | Imp A1 A2 => Imp (subst k A1 d) (subst k A2 d)
+ | Atom P ld =>
+ Atom P (subst_list ld)
+ end.
+
+Notation "A ^^ d" := (subst 0 A d) (at level 67).
+Notation "A ^ x" := (subst 0 A (fvar x)).
+
+(** A formula is [locl] ("locally closed") when all [bvar]s are bound
+ by quantifiers (while there might be [fvar]s around) *)
+Definition locl (A:formula) := forall k d, (subst k A d) = A.
+Definition locli (d':indiv) := forall k d, (subst_indiv k d' d) = d'.
+
+Definition context := list (var*formula).
+Hint Unfold context.
+
+(** Natural deduction system with cofinite quantification *)
+Inductive proof : context -> term -> formula -> Set :=
+| DisjIL Gamma t A B : proof Gamma t A -> proof Gamma (Inl t) (Disj A B)
+
+| DisjIR Gamma t A B : proof Gamma t B -> proof Gamma (Inr t) (Disj A B)
+
+| ImpI Gamma a t A B :
+ proof ((a,A)::Gamma) t B -> proof Gamma (Lam a t) (Imp A B)
+
+| Hypo Gamma a A : In (a,A) Gamma -> proof Gamma (Var a) A
+
+| DisjE Gamma a b e t u A B C :
+ proof Gamma e (Disj A B) ->
+ proof ((a,A)::Gamma) t C -> proof ((b,B)::Gamma) u C ->
+ proof Gamma (Match e a t b u) C
+
+| ImpE Gamma t u A B : proof Gamma t (Imp A B) ->
+ proof Gamma u A -> proof Gamma (App t u) B
+
+| AllI Gamma t A L : locl (All A) ->
+ (forall x, x \notin L -> proof Gamma t (A^x)) ->
+ proof Gamma (QLam t) (All A)
+
+| AllE Gamma t A d : locli d ->
+ proof Gamma t (All A) -> proof Gamma (QApp t) (A^^d)
+
+| ExI Gamma t A d : locli d ->
+ proof Gamma t (A^^d) -> proof Gamma (QWit t) (Ex A)
+
+| ExE Gamma a e t A C L : locl (Ex A) ->
+ proof Gamma e (Ex A) ->
+ (forall x, x\notin L -> proof ((a,A^x)::Gamma) t C) ->
+ proof Gamma (QMatch e a t) C
+ .
+
+Section Abstract_Semantics.
+ (** An abstract Kripke-style structure: [wle] is the preorder, [X] is
+ the exploding-world predicate, [aforces] is strong forcing of
+ atomic formulae. *)
+ Record Kripke : Type := {
+ world :> Set;
+ wle : world -> world -> Type;
+ wle_refl : forall w, wle w w;
+ wle_trans : forall w w' w'', wle w w' -> wle w' w'' -> wle w w'';
+ X : world -> formula -> Set;
+ aforces : world -> predicate -> list indiv -> Set;
+ aforces_mon : forall w P ld, @aforces w P ld -> forall w', wle w w' -> @aforces w' P ld
+ }.
+ Notation "w <= w'" := (wle w w').
+
+ Variable K:Kripke.
+
+ (** The continuations monad we will use to extend forcing to
+ composite formulas *)
+ Definition Kont (F:K->formula->Type)(w:K)(A:formula) :=
+ forall C:formula,
+ (forall w1, w <= w1 ->
+ (forall w2, w1 <= w2 -> F w2 A -> X w2 C) -> X w1 C).
+ Hint Unfold Kont.
+
+ (** Strong forcing extended to composite formulas. The expression
+ [sforces w A] stands for strong forcing, while [Kont sforces w A]
+ stands for (weak) forcing of A in the world w. *)
+ Fixpoint sforces' (n:nat)(w:K)(A:formula) {struct n} : Type :=
+ match n with
+ | O => Empty_set
+ | S m =>
+ match A with
+ | Atom P ld => aforces w P ld
+ | Imp A1 A2 => forall w', w <= w'-> Kont (sforces' m) w' A1 -> Kont (sforces' m) w' A2
+ | Disj A1 A2 => Kont (sforces' m) w A1 + Kont (sforces' m) w A2
+ | All A1 => forall w', w <= w' -> forall d, locli d -> Kont (sforces' m) w' (A1^^d)
+ | Ex A1 => {d:indiv & (locli d) * Kont (sforces' m) w (A1^^d)}
+ end
+ end.
+
+ Fixpoint depth (A:formula) : nat :=
+ match A with
+ | Atom _ _ => 1
+ | Imp A1 A2 => S (max (depth A1) (depth A2))
+ | Disj A1 A2 => S (max (depth A1) (depth A2))
+ | All A1 => S (depth A1)
+ | Ex A1 => S (depth A1)
+ end.
+
+ (** However, we can not define strong forcing as simply as when we had no quantifiers, because we need to do a recursion on the _complexity_ of a formula, not just its structure; we do that using the measure of [depth] above. *)
+ Definition sforces (w:K)(A:formula) := sforces' (depth A) w A.
+ Hint Unfold sforces.
+
+ (** And now this section proves that the definition using the measure is correct *)
+ Section sforces_correctness.
+ Lemma depth_subst : forall A d k, depth (subst k A d) = depth A.
+ Proof.
+ induction A; simpl; intros.
+
+ reflexivity.
+
+ rewrite <- IHA1 with d k.
+ rewrite <- IHA2 with d k.
+ reflexivity.
+
+ rewrite <- IHA1 with d k.
+ rewrite <- IHA2 with d k.
+ reflexivity.
+
+ rewrite <- IHA with d (S k).
+ reflexivity.
+
+ rewrite <- IHA with d (S k).
+ reflexivity.
+ Defined.
+
+ Lemma sforces'_S : forall n A w, le (depth A) n ->
+ (sforces' n w A <-> sforces' (S n) w A).
+ Proof.
+ induction n;
+ intros A w Hle.
+
+ inversion Hle.
+ destruct A;
+ inversion H0.
+
+ split.
+
+ (* Direction ==> *)
+ intro H.
+ simpl in H.
+ hnf.
+ assert (IHn1 := fun A w Hle => fst (IHn A w Hle)).
+ assert (IHn2 := fun A w Hle => snd (IHn A w Hle)).
+ clear IHn.
+ destruct A.
+
+ assumption.
+
+ intros w' ww' HKA1Sn.
+ intros C1 w1 w'w1 k.
+ apply H with w'.
+ assumption.
+ intros C2 w2 w1w2 HA1n.
+ apply HKA1Sn.
+ assumption.
+ intros w3 w2w3 HA1Sn.
+ apply HA1n.
+ assumption.
+ apply IHn2.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_l.
+ apply Hle.
+ assumption.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn1.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_r.
+ apply Hle.
+ assumption.
+
+ case H; intro H'.
+ left.
+ intros C1 w1 ww1 k.
+ apply H'.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn1; auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_l.
+ apply Hle.
+ right.
+ intros C1 w1 ww1 k.
+ apply H'.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn1; auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_r.
+ apply Hle.
+
+ intros.
+ intros C1 w1 w'w1 k.
+ apply H with w1 d.
+ eauto using wle_trans.
+ assumption.
+ apply wle_refl.
+ intros w2 w1w2 HA.
+ apply k.
+ assumption.
+ apply IHn1;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ rewrite depth_subst.
+ assumption.
+
+ destruct H as [d [H1 H2]].
+ exists d.
+ split; [assumption | ].
+ intros C1 w1 ww1 k.
+ apply H2.
+ assumption.
+ intros w2 w1w2 H3.
+ apply k.
+ assumption.
+ apply IHn1;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ rewrite depth_subst.
+ assumption.
+
+
+ (* Direction <== (exact copy-past of the other direction except for
+ swapping IHn1 <~> IHn2) *)
+ intro H.
+ simpl.
+ hnf in H.
+ assert (IHn1 := fun A w Hle => fst (IHn A w Hle)).
+ assert (IHn2 := fun A w Hle => snd (IHn A w Hle)).
+ clear IHn.
+ destruct A.
+
+ assumption.
+
+ intros w' ww' HKA1Sn.
+ intros C1 w1 w'w1 k.
+ apply H with w'.
+ assumption.
+ intros C2 w2 w1w2 HA1n.
+ apply HKA1Sn.
+ assumption.
+ intros w3 w2w3 HA1Sn.
+ apply HA1n.
+ assumption.
+ apply IHn1.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_l.
+ apply Hle.
+ assumption.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn2.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_r.
+ apply Hle.
+ assumption.
+
+ case H; intro H'.
+ left.
+ intros C1 w1 ww1 k.
+ apply H'.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn2; auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_l.
+ apply Hle.
+ right.
+ intros C1 w1 ww1 k.
+ apply H'.
+ assumption.
+ intros.
+ apply k.
+ assumption.
+ apply IHn2; auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ eapply le_trans.
+ apply le_max_r.
+ apply Hle.
+
+ intros.
+ intros C1 w1 w'w1 k.
+ apply H with w1 d.
+ eauto using wle_trans.
+ assumption.
+ apply wle_refl.
+ intros w2 w1w2 HA.
+ apply k.
+ assumption.
+ apply IHn2;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ rewrite depth_subst.
+ assumption.
+
+ destruct H as [d [H1 H2]].
+ exists d.
+ split; [assumption | ].
+ intros C1 w1 ww1 k.
+ apply H2.
+ assumption.
+ intros w2 w1w2 H3.
+ apply k.
+ assumption.
+ apply IHn2;auto.
+ clear -Hle.
+ simpl in *.
+ apply le_S_n in Hle.
+ rewrite depth_subst.
+ assumption.
+ Defined.
+
+ Definition sforces'_S1 := fun n A w Hle => fst (@sforces'_S n A w Hle).
+ Definition sforces'_S2 := fun n A w Hle => snd (@sforces'_S n A w Hle).
+
+ Lemma sforces'_mon1 : forall n m, le n m ->
+ forall A, le (depth A) n -> forall w, sforces' n w A -> sforces' m w A.
+ Proof.
+ intros n m Hle.
+ induction Hle.
+
+ intros.
+ auto.
+
+ intros.
+ apply sforces'_S1.
+ eauto using le_trans.
+ auto.
+ Defined.
+
+ Lemma sforces'_mon2 : forall n m, le n m ->
+ forall A, le (depth A) n -> forall w, sforces' m w A -> sforces' n w A.
+ Proof.
+ intros n m Hle.
+ induction Hle.
+
+ intros.
+ auto.
+
+ intros.
+ apply IHHle.
+ assumption.
+ apply sforces'_S2.
+ eauto using le_trans.
+ assumption.
+ Defined.
+
+ Lemma Kont_invar : forall F1 F2:K->formula->Type, forall A,
+ (forall w, F1 w A -> F2 w A) -> forall w, Kont F1 w A -> Kont F2 w A.
+ Proof.
+ intros F1 F2 A HF1F2.
+ intros w HF1.
+ intros C1 w1 ww1 k1.
+ apply HF1.
+ assumption.
+ intros w2 w1w2 HF1'.
+ apply k1.
+ assumption.
+ auto.
+ Defined.
+
+ Lemma sforces_correct_Atom : forall w P ld,
+ sforces w (@Atom P ld) -> aforces w P ld.
+ Proof.
+ unfold sforces.
+ simpl.
+ auto.
+ Defined.
+
+ Lemma sforces_correct_Atom' : forall w P ld,
+ aforces w P ld -> sforces w (@Atom P ld).
+ Proof.
+ unfold sforces.
+ simpl.
+ auto.
+ Defined.
+
+ Lemma sforces_correct_Disj : forall w A B, sforces w (Disj A B) ->
+ Kont sforces w A + Kont sforces w B.
+ Proof.
+ intros w A B H.
+ unfold sforces in H.
+ simpl in H.
+ destruct H.
+
+ left.
+ change (Kont (sforces' (depth A)) w A).
+ generalize dependent k.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon2.
+ apply le_max_l.
+ constructor.
+
+ right.
+ change (Kont (sforces' (depth B)) w B).
+ generalize dependent k.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon2.
+ apply le_max_r.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_Disj' : forall w A B,
+ Kont sforces w A + Kont sforces w B -> sforces w (Disj A B).
+ Proof.
+ intros w A B H.
+ unfold sforces.
+ simpl.
+ destruct H.
+
+ left.
+ generalize dependent k.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon1.
+ apply le_max_l.
+ constructor.
+
+ right.
+ generalize dependent k.
+ apply Kont_invar.
+ intros w'.
+ apply sforces'_mon1.
+ apply le_max_r.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_Imp : forall w A B, sforces w (Imp A B) ->
+ forall w', w <= w' -> Kont sforces w' A -> Kont sforces w' B.
+ Proof.
+ intros w A B H.
+ unfold sforces in H.
+ simpl in H.
+ intros w' ww' HA.
+
+ change (Kont (sforces' (depth A)) w' A) in HA.
+ change (Kont (sforces' (depth B)) w' B).
+
+ apply Kont_invar with (sforces' (max (depth A) (depth B))).
+ apply sforces'_mon2.
+ apply le_max_r.
+ constructor.
+ apply H.
+ assumption.
+ generalize dependent HA.
+ apply Kont_invar.
+ apply sforces'_mon1.
+ apply le_max_l.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_Imp' : forall w A B,
+ (forall w', w <= w' -> Kont sforces w' A -> Kont sforces w' B)
+ -> sforces w (Imp A B).
+ Proof.
+ intros w A B H.
+ unfold sforces.
+ simpl.
+ intros w' ww' HA.
+
+ apply Kont_invar with (sforces' (depth B)).
+ intros w''.
+ apply sforces'_mon1.
+ apply le_max_r.
+ constructor.
+ apply H.
+ assumption.
+ generalize dependent HA.
+ apply Kont_invar.
+ intros w''.
+ apply sforces'_mon2.
+ apply le_max_l.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_All : forall w A, sforces w (All A) ->
+ forall w', w <= w' -> forall d, locli d -> Kont sforces w' (A^^d).
+ Proof.
+ intros w A H.
+ unfold sforces in H.
+ simpl in H.
+ intros w' ww' d Hd.
+
+ change (Kont (sforces' (depth (A^^d))) w' (A^^d)).
+ assert (H' := H w' ww' d Hd).
+ generalize dependent H'.
+ apply Kont_invar.
+ apply sforces'_mon2.
+ rewrite depth_subst.
+ constructor.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_All' : forall w A,
+ (forall w', w <= w' -> forall d, locli d -> Kont sforces w' (A^^d))
+ -> sforces w (All A).
+ Proof.
+ intros w A H.
+ unfold sforces.
+ simpl.
+ intros w' ww' d Hd.
+ rewrite <- depth_subst with A d 0.
+ apply H.
+ assumption.
+ assumption.
+ Defined.
+
+ Lemma sforces_correct_Ex : forall w A, sforces w (Ex A) ->
+ sigT (fun d => (locli d) * Kont sforces w (A^^d)).
+ Proof.
+ intros w A H.
+ unfold sforces in H.
+ simpl in H.
+ destruct H as [d [Hd H]].
+
+ exists d.
+ split; [assumption|].
+ change (Kont (sforces' (depth (A^^d))) w (A^^d)).
+ generalize dependent H.
+ apply Kont_invar.
+ apply sforces'_mon2.
+ rewrite depth_subst.
+ constructor.
+ constructor.
+ Defined.
+
+ Lemma sforces_correct_Ex' : forall w A,
+ (sigT (fun d => (locli d) * Kont sforces w (A^^d)))
+ -> sforces w (Ex A).
+ Proof.
+ intros w A H.
+ unfold sforces.
+ simpl.
+ destruct H as [d [Hd H]].
+ exists d.
+ split; [assumption|].
+ rewrite <- depth_subst with A d 0.
+ apply H.
+ Defined.
+ End sforces_correctness.
+
+ Notation "w ||+ A " := (sforces w A) (at level 70).
+ Notation "w ||- A " := (Kont sforces w A) (at level 70).
+
+ Fixpoint forces_cxt (w:K)(Gamma:context) : Type :=
+ match Gamma with
+ | nil => unit
+ | cons aA Gamma' => prod (w ||- (snd aA)) (forces_cxt w Gamma')
+ end.
+
+ Notation " w ||-- Gamma" := (forces_cxt w Gamma) (at level 70).
+
+ Lemma sforces_mon : forall A w, w ||+ A -> forall w', w <= w' -> w' ||+ A .
+ Proof.
+ induction A.
+
+ unfold sforces.
+ simpl.
+ intros.
+ eauto using aforces_mon.
+
+ unfold sforces.
+ intros w H1 w' ww'.
+ simpl in *.
+ eauto using wle_trans.
+
+ unfold sforces.
+ intros w H1 w' ww'.
+ simpl in *.
+ case H1; intro H'.
+ left.
+ eauto using wle_trans.
+ eauto using wle_trans.
+
+ unfold sforces.
+ intros w H1 w' ww'.
+ simpl in *.
+ eauto using wle_trans.
+
+ unfold sforces.
+ intros w H1 w' ww'.
+ simpl in *.
+ destruct H1 as [d [Hd H]].
+ exists d.
+ split; [assumption|].
+ eauto using wle_trans.
+ Defined.
+
+ Definition ret {w A} : w ||+ A -> w ||- A.
+ Proof.
+ intros H.
+ intros C w1 w_w1 k.
+ apply k.
+ apply wle_refl.
+ apply sforces_mon with w.
+ assumption.
+ assumption.
+ Defined.
+
+ Definition bind {w A B} : (forall w', wle w w' -> w' ||+ A -> w' ||- B) -> w ||- A -> w ||- B.
+ Proof.
+ intros f a.
+ intros C w1 ww1 k.
+ apply a.
+ assumption.
+ intros w2 w1w2 a'.
+ apply (f w2).
+ eauto using wle_trans.
+ exact a'.
+ apply wle_refl.
+ eauto using wle_trans.
+ Defined.
+
+ Lemma forces_mon : forall A w, w ||- A -> forall w', w <= w' -> w' ||- A.
+ Proof.
+ induction A;
+ intros;
+ eauto using wle_trans.
+ Defined.
+
+ Lemma forces_cxt_mon : forall Gamma w, w ||-- Gamma -> forall w', w <= w' -> w' ||-- Gamma.
+ Proof.
+ induction Gamma.
+ simpl.
+ auto.
+ simpl.
+ intros.
+ destruct X0.
+ split; eauto using wle_trans,forces_mon.
+ Defined.
+
+ Lemma forces_cxt_In : forall w x A Gamma, In (x, A) Gamma -> w ||-- Gamma -> w ||- A.
+ Proof.
+ induction Gamma.
+ simpl.
+ intros.
+ contradict H.
+ simpl.
+ intros H H0.
+ destruct H.
+ rewrite e in H0.
+ intuition.
+ intuition.
+ Defined.
+End Abstract_Semantics.
+
+(** Various lemmas that have to do with substitutions related to quantifier bound variables; this was the hardest part of the formalisation. *)
+Section subst_lemmas.
+ Lemma locli_fvar : forall x, locli (fvar x).
+ Proof.
+ unfold locli.
+ simpl.
+ reflexivity.
+ Defined.
+
+ Lemma func_wrapper1 : forall f l l', func f l = func f l' -> l = l'.
+ Proof.
+ intros.
+ congruence.
+ Defined.
+
+ Lemma func_wrapper2 : forall f l l', l = l' -> func f l = func f l'.
+ Proof.
+ intros.
+ congruence.
+ Defined.
+
+ Lemma Atom_wrapper1 : forall f l l', Atom f l = Atom f l' -> l = l'.
+ Proof.
+ intros.
+ congruence.
+ Defined.
+
+ Lemma Atom_wrapper2 : forall p l l', l = l' -> Atom p l = Atom p l'.
+ Proof.
+ intros.
+ congruence.
+ Defined.
+
+ Lemma cons_eqs : forall T a a' l l', @cons T a l = @cons T a' l' ->
+ (a=a') * (l=l').
+ Proof.
+ intros.
+ inversion H.
+ auto.
+ Defined.
+
+ Lemma eqs_cons : forall T a a' l l',
+ (a=a') * (l=l') -> @cons T a l = @cons T a' l'.
+ Proof.
+ intros.
+ destruct H.
+ rewrite e.
+ rewrite e0.
+ auto.
+ Defined.
+
+ Lemma subst_indiv_zero_twice : forall a d d' n', locli d ->
+ subst_indiv n' (subst_indiv n' a d) d' = subst_indiv n' a d.
+ Proof.
+ fix 1.
+ intro a.
+ destruct a.
+
+ intros.
+ simpl.
+ assert (H' := beq_nat_false n' n).
+ destruct beq_nat.
+ apply H.
+ simpl.
+ assert (H'' := beq_nat_true n' n).
+ destruct beq_nat.
+ assert (False).
+ apply H'.
+ reflexivity.
+ apply H''.
+ reflexivity.
+ intuition.
+ reflexivity.
+
+ simpl.
+ reflexivity.
+
+ simpl.
+ intros.
+ apply func_wrapper2.
+ induction l.
+ reflexivity.
+ apply eqs_cons.
+ split.
+ apply subst_indiv_zero_twice.
+ assumption.
+
+ apply IHl.
+ Defined.
+
+ Lemma locli_locli_subst : forall a m k d c,
+ (forall k c, subst_indiv (S (k+m)) a c = a) ->
+ (forall k c, subst_indiv (k+m) d c = d) ->
+ subst_indiv (k+m) (subst_indiv m a d) c = subst_indiv m a d.
+ Proof.
+ fix 1.
+ destruct a.
+
+ intros.
+ simpl.
+ remember (beq_nat m n).
+ destruct b.
+ apply H0.
+ destruct k.
+ simpl.
+ rewrite <- Heqb.
+ reflexivity.
+ rewrite plus_Sn_m.
+ apply H.
+
+ simpl.
+ reflexivity.
+
+ intros.
+ simpl.
+ simpl in H.
+ apply func_wrapper2.
+ assert (H' := fun k d => func_wrapper1 (H k d)).
+ clear H.
+ induction l.
+ reflexivity.
+ apply eqs_cons.
+ assert (H1 := fun k d => fst (cons_eqs (H' k d))).
+ assert (H2 := fun k d => snd (cons_eqs (H' k d))).
+ clear H'.
+ split.
+ rewrite locli_locli_subst.
+ reflexivity.
+ apply H1.
+ assumption.
+ apply IHl.
+ apply H2.
+ Defined.
+
+ Lemma locl_locli_subst : forall A m,
+ (forall k d, subst (S (k + m)) A d = A) ->
+ forall d,
+ (forall k d0, subst_indiv (k + m) d d0 = d) ->
+ forall k' d',
+ subst (k'+m) (subst m A d) d' = subst m A d.
+ Proof.
+ fix 1.
+ intros.
+ rename H into H'.
+ destruct A.
+
+ simpl in *.
+ apply Atom_wrapper2.
+ assert (H'' := fun k d => Atom_wrapper1 (H' k d)).
+ clear H'.
+ induction l.
+ reflexivity.
+ assert (H1 := fun k d => fst (cons_eqs (H'' k d))).
+ assert (H2 := fun k d => snd (cons_eqs (H'' k d))).
+ clear H''.
+ apply eqs_cons.
+ split.
+ rewrite locli_locli_subst.
+ reflexivity.
+ apply H1.
+ apply H0.
+ apply IHl.
+ apply H2.
+
+ simpl in *.
+ rewrite locl_locli_subst.
+ rewrite locl_locli_subst.
+ reflexivity.
+ unfold locl.
+ simpl.
+ intros.
+ assert (H'' := H' k d0).
+ congruence.
+ assumption.
+ unfold locl.
+ simpl.
+ intros.
+ assert (H'' := H' k d0).
+ congruence.
+ assumption.
+
+ simpl in *.
+ rewrite locl_locli_subst.
+ rewrite locl_locli_subst.
+ reflexivity.
+ unfold locl.
+ simpl.
+ intros.
+ assert (H'' := H' k d0).
+ congruence.
+ assumption.
+ unfold locl.
+ simpl.
+ intros.
+ assert (H'' := H' k d0).
+ congruence.
+ assumption.
+
+ simpl in *.
+ rewrite plus_n_Sm.
+ rewrite locl_locli_subst.
+ reflexivity.
+ intros.
+ assert (H'' := H' k d0).
+ rewrite <- plus_n_Sm.
+ congruence.
+ intros.
+ rewrite <- Arith.Plus.plus_Snm_nSm.
+ apply H0.
+
+ simpl in *.
+ rewrite plus_n_Sm.
+ rewrite locl_locli_subst.
+ reflexivity.
+ intros.
+ assert (H'' := H' k d0).
+ rewrite <- plus_n_Sm.
+ congruence.
+ intros.
+ rewrite <- Arith.Plus.plus_Snm_nSm.
+ apply H0.
+ Defined.
+
+ Lemma locl_locli_subst' : forall A, locl (All A) ->
+ forall d, locli d -> locl (A^^d).
+ Proof.
+ unfold locl,locli.
+ simpl.
+ intros.
+ replace k with (plus k 0).
+ apply locl_locli_subst.
+ intros.
+ assert (H' := H (plus k0 0) d1).
+ congruence.
+ intros.
+ apply H0.
+ apply Arith.Plus.plus_0_r.
+ Defined.
+
+ Lemma locl_locli_subst'' : forall A, locl (Ex A) ->
+ forall d, locli d -> locl (A^^d).
+ Proof.
+ unfold locl,locli.
+ simpl.
+ intros.
+ replace k with (plus k 0).
+ apply locl_locli_subst.
+ intros.
+ assert (H' := H (plus k0 0) d1).
+ congruence.
+ intros.
+ apply H0.
+ apply Arith.Plus.plus_0_r.
+ Defined.
+
+ Fixpoint lsubst (T:Type)(k:nat)(l:list (T*formula))(d:indiv) :=
+ match l with
+ | cons aA aAs => cons (fst aA, (subst k (snd aA) d)) (lsubst k aAs d)
+ | nil => nil
+ end.
+
+ Notation "l \ x" := (lsubst 0 l (fvar x)) (at level 60).
+
+ Fixpoint fvar_swap_indiv (x:var_free)(a:indiv)(y:var_free) {struct a} : indiv :=
+ let fvar_swap_list := fix fvar_swap_list
+ (l:list indiv) {struct l} : list indiv :=
+ match l with
+ | nil => nil
+ | cons a' l' => cons (fvar_swap_indiv x a' y) (fvar_swap_list l')
+ end in
+ match a with
+ | func f ld =>
+ func f (fvar_swap_list ld)
+ | fvar z => if (var_free_dec x z) then (fvar y) else (fvar z)
+ | bvar i => bvar i
+ end.
+
+ Fixpoint fvar_swap_typ (x:var_free)(A:formula)(y:var_free) {struct A} : formula :=
+ let fvar_swap_list := fix fvar_swap_list
+ (l:list indiv) {struct l} : list indiv :=
+ match l with
+ | nil => nil
+ | cons a' l' => cons (fvar_swap_indiv x a' y) (fvar_swap_list l')
+ end in
+ match A with
+ | Ex A1 => Ex (fvar_swap_typ x A1 y)
+ | All A1 => All (fvar_swap_typ x A1 y)
+ | Disj A1 A2 => Disj (fvar_swap_typ x A1 y) (fvar_swap_typ x A2 y)
+ | Imp A1 A2 => Imp (fvar_swap_typ x A1 y) (fvar_swap_typ x A2 y)
+ | Atom P ld =>
+ Atom P (fvar_swap_list ld)
+ end.
+
+ Fixpoint fvar_swap_cxt (T:Type)(x:var_free)(l:list (T*formula))(y:var_free) :=
+ match l with
+ | cons aA aAs =>
+ cons (fst aA, (fvar_swap_typ x (snd aA) y)) (fvar_swap_cxt x aAs y)
+ | nil => nil
+ end.
+
+ Lemma locli_fvar_swap : forall a m,
+ (forall k d, subst_indiv (plus k m) a d = a) ->
+ forall x y k d,
+ subst_indiv (plus k m) (fvar_swap_indiv x a y) d = fvar_swap_indiv x a y.
+ Proof.
+ fix 1.
+ intro a.
+ destruct a.
+
+ simpl.
+ auto.
+
+ simpl.
+ intros.
+ destruct var_free_dec.
+ Hint Unfold locli_fvar.
+ apply locli_fvar.
+ apply locli_fvar.
+
+ intros.
+ simpl in *.
+ assert (H' := fun k d => func_wrapper1 (H k d)).
+ clear H.
+ apply func_wrapper2.
+ induction l.
+ reflexivity.
+ apply eqs_cons.
+ assert (H1 := fun k d => fst (cons_eqs (H' k d))).
+ assert (H2 := fun k d => snd (cons_eqs (H' k d))).
+ clear H'.
+ split.
+ unfold locli in locli_fvar_swap.
+ rewrite locli_fvar_swap.
+ reflexivity.
+ apply H1.
+ apply IHl.
+ apply H2.
+ Defined.
+
+ Lemma locl_fvar_swap : forall A m,
+ (forall k d, subst (plus k m) A d = A) ->
+ forall x y k d,
+ subst (plus k m) (fvar_swap_typ x A y) d = fvar_swap_typ x A y.
+ Proof.
+ induction A.
+
+ induction l.
+ simpl.
+ auto.
+ intros m H.
+ simpl in H.
+ assert (H' := fun k d => cons_eqs (Atom_wrapper1 (H k d))).
+ clear H.
+ assert (H1 := fun k d => fst (H' k d)).
+ assert (H2 := fun k d => snd (H' k d)).
+ clear H'.
+ intros.
+ simpl.
+ unfold locl.
+ intros.
+ simpl.
+ apply Atom_wrapper2.
+ apply eqs_cons.
+ split.
+ apply locli_fvar_swap.
+ apply H1.
+ simpl in IHl.
+ assert (IHl' := fun m H x y k d => Atom_wrapper1 (IHl m H x y k d)).
+ clear IHl.
+ assert (H2' : (forall (k : nat) (d : indiv),
+ Atom p
+ ((fix subst_list (l : list indiv) : list indiv :=
+ match l with
+ | nil => nil
+ | a' :: l' => subst_indiv (plus k m) a' d :: subst_list l'
+ end) l) = Atom p l)).
+ intros.
+ assert (H2'' := H2 k0 d0).
+ congruence.
+ apply IHl' with m x y k d in H2'.
+ clear IHl'.
+ apply H2'.
+
+ unfold locl.
+ simpl.
+ intros.
+ rewrite IHA1.
+ rewrite IHA2.
+ reflexivity.
+ assert (H' := fun k d => H k d).
+ intros k' d'.
+ assert (H'' := H' k' d').
+ congruence.
+ assert (H' := fun k d => H k d).
+ intros k' d'.
+ assert (H'' := H' k' d').
+ congruence.
+
+ unfold locl.
+ simpl.
+ intros.
+ rewrite IHA1.
+ rewrite IHA2.
+ reflexivity.
+ assert (H' := fun k d => H k d).
+ intros k' d'.
+ assert (H'' := H' k' d').
+ congruence.
+ assert (H' := fun k d => H k d).
+ intros k' d'.
+ assert (H'' := H' k' d').
+ congruence.
+
+ unfold locl.
+ simpl.
+ intros.
+ rewrite plus_n_Sm.
+ rewrite IHA.
+ reflexivity.
+ intros k' d'.
+ rewrite <- plus_n_Sm.
+ assert (H'' := H k' d').
+ congruence.
+
+ unfold locl.
+ simpl.
+ intros.
+ rewrite plus_n_Sm.
+ rewrite IHA.
+ reflexivity.
+ intros k' d'.
+ rewrite <- plus_n_Sm.
+ assert (H'' := H k' d').
+ congruence.
+ Defined.
+
+ Lemma subst_fvar_swap_indiv : forall i k d x y,
+ fvar_swap_indiv x (subst_indiv k i d) y
+ = subst_indiv k (fvar_swap_indiv x i y) (fvar_swap_indiv x d y).
+ Proof.
+ fix 1.
+ intro i.
+ destruct i.
+
+ simpl.
+ intros.
+ destruct beq_nat.
+ reflexivity.
+ reflexivity.
+
+ simpl.
+ intros.
+ destruct var_free_dec.
+ reflexivity.
+ reflexivity.
+
+ intros.
+ simpl.
+ apply func_wrapper2.
+ induction l.
+ reflexivity.
+ apply eqs_cons.
+ split.
+ rewrite subst_fvar_swap_indiv.
+ reflexivity.
+ apply IHl.
+ Defined.
+
+ Lemma subst_fvar_swap : forall A k d x y, fvar_swap_typ x (subst k A d) y
+ = subst k (fvar_swap_typ x A y) (fvar_swap_indiv x d y).
+ Proof.
+ fix 1.
+ intro A.
+ destruct A.
+
+ intros.
+ simpl.
+
+ apply Atom_wrapper2.
+ induction l.
+ reflexivity.
+ apply eqs_cons.
+ split.
+ rewrite subst_fvar_swap_indiv.
+ reflexivity.
+ apply IHl.
+
+ simpl.
+ intros.
+ rewrite subst_fvar_swap.
+ rewrite subst_fvar_swap.
+ reflexivity.
+
+ simpl.
+ intros.
+ rewrite subst_fvar_swap.
+ rewrite subst_fvar_swap.
+ reflexivity.
+
+ simpl.
+ intros.
+ rewrite subst_fvar_swap.
+ reflexivity.
+
+ simpl.
+ intros.
+ rewrite subst_fvar_swap.
+ reflexivity.
+ Defined.
+
+ Fixpoint fvar_appears_indiv (x:var_free)(a:indiv) {struct a} : bool :=
+ let fvar_appears_list :=
+ fix fvar_appears_list (x:var_free)(l:list indiv) {struct l} : bool :=
+ match l with
+ | nil => false
+ | cons d l' =>
+ orb (fvar_appears_indiv x d) (fvar_appears_list x l')
+ end
+ in match a with
+ | func f ld => fvar_appears_list x ld
+ | fvar z => if (var_free_dec x z) then true else false
+ | bvar i => false
+ end.
+
+ Fixpoint fvar_appears_typ (x:var_free)(A:formula) {struct A} : bool :=
+ let fvar_appears_list :=
+ fix fvar_appears_list (x:var_free)(l:list indiv) {struct l} : bool :=
+ match l with
+ | nil => false
+ | cons d l' =>
+ orb (fvar_appears_indiv x d) (fvar_appears_list x l')
+ end
+ in match A with
+ | Ex A1 => fvar_appears_typ x A1
+ | All A1 => fvar_appears_typ x A1
+ | Disj A1 A2 => orb (fvar_appears_typ x A1) (fvar_appears_typ x A2)
+ | Imp A1 A2 => orb (fvar_appears_typ x A1) (fvar_appears_typ x A2)
+ | Atom P ld => fvar_appears_list x ld
+ end.
+
+ Fixpoint fvar_appears_cxt (T:Type)(x:var_free)(l:list (T*formula)) : bool :=
+ match l with
+ | cons aA aAs =>
+ orb (fvar_appears_typ x (snd aA)) (fvar_appears_cxt x aAs)
+ | nil => false
+ end.
+
+ Lemma swap_appears_not_indiv : forall d x,
+ fvar_appears_indiv x d = false -> forall y, fvar_swap_indiv x d y = d.
+ Proof.
+ fix 1.
+ intro d.
+ destruct d.
+
+ simpl.
+ auto.
+
+ simpl.
+ intros.
+ destruct var_free_dec.
+ discriminate.
+ reflexivity.
+
+ intros.
+ simpl in *.
+ apply func_wrapper2.
+ induction l.
+ reflexivity.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ apply eqs_cons.
+ split.
+ apply swap_appears_not_indiv.
+ assumption.
+ apply IHl.
+ assumption.
+ Defined.
+
+ Lemma swap_appears_not_typ : forall A x,
+ fvar_appears_typ x A = false -> forall y, fvar_swap_typ x A y = A.
+ Proof.
+ fix 1.
+ intro A.
+ destruct A.
+
+ intros.
+ simpl in *.
+ apply Atom_wrapper2.
+ induction l.
+ reflexivity.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ apply eqs_cons.
+ split.
+ apply swap_appears_not_indiv.
+ assumption.
+ apply IHl.
+ assumption.
+
+ simpl.
+ intros.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ rewrite swap_appears_not_typ.
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ assumption.
+ assumption.
+
+ simpl.
+ intros.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ rewrite swap_appears_not_typ.
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ assumption.
+ assumption.
+
+ simpl.
+ intros.
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ assumption.
+
+ simpl.
+ intros.
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ assumption.
+ Defined.
+
+ Lemma swap_appears_not_cxt : forall T, forall GD:list (T*formula), forall x,
+ fvar_appears_cxt x GD = false -> forall y, fvar_swap_cxt x GD y = GD.
+ Proof.
+ induction GD.
+
+ reflexivity.
+
+ simpl.
+ intros.
+ apply Bool.orb_false_elim in H.
+ destruct H as [H1 H2].
+ rewrite swap_appears_not_typ; auto.
+ rewrite IHGD;auto.
+ destruct a; auto.
+ Defined.
+
+ Fixpoint FV_indiv (d:indiv) {struct d} : list var_free :=
+ let FV_indiv_list :=
+ fix FV_indiv_list (l:list indiv) {struct l} : list var_free :=
+ match l with
+ | nil => nil
+ | cons d' l' => List.app (FV_indiv d') (FV_indiv_list l')
+ end in
+ match d with
+ | func f vl => FV_indiv_list vl
+ | fvar z => (z::nil)
+ | bvar i => nil
+ end.
+
+ Fixpoint FV_typ (A:formula) {struct A} : list var_free :=
+ let FV_indiv_list :=
+ fix FV_indiv_list (l:list indiv) {struct l} : list var_free :=
+ match l with
+ | nil => nil
+ | cons d' l' => List.app (FV_indiv d') (FV_indiv_list l')
+ end in
+ match A with
+ | Ex A1 => FV_typ A1
+ | All A1 => FV_typ A1
+ | Disj A1 A2 => (FV_typ A1) ++ (FV_typ A2)
+ | Imp A1 A2 => (FV_typ A1) ++ (FV_typ A2)
+ | Atom P vl => FV_indiv_list vl
+ end.
+
+ Fixpoint FV_cxt (T:Type)(l:list (T*formula)) {struct l} : list var_free :=
+ match l with
+ | cons aA aAs => FV_typ (snd aA) ++ (FV_cxt aAs)
+ | nil => nil
+ end.
+
+ Lemma notin_app : forall T:Type, forall x:T, forall L1 L2,
+ x\notin L1++L2 -> (x\notin L1)*(x\notin L2).
+ Proof.
+ intros.
+ split.
+ eauto using in_or_app.
+ eauto using in_or_app.
+ Defined.
+
+ Lemma notin_FV_not_appears_indiv : forall a x, x \notin FV_indiv a ->
+ fvar_appears_indiv x a = false.
+ Proof.
+ fix 1.
+ destruct a.
+
+ simpl.
+ reflexivity.
+
+ simpl.
+ intros.
+ destruct var_free_dec.
+ symmetry in e.
+ intuition.
+ reflexivity.
+
+ intros.
+ induction l.
+ reflexivity.
+ simpl in *.
+ apply Bool.orb_false_intro.
+ apply notin_app in H.
+ destruct H.
+ apply notin_FV_not_appears_indiv.
+ assumption.
+ apply IHl.
+ apply notin_app in H.
+ destruct H.
+ assumption.
+ Defined.
+
+ Lemma notin_FV_not_appears : forall A x, x \notin FV_typ A ->
+ fvar_appears_typ x A = false.
+ Proof.
+ induction A.
+
+ simpl.
+ induction l.
+ reflexivity.
+ intros.
+ apply Bool.orb_false_intro.
+ apply notin_app in H.
+ destruct H as [H _].
+ apply notin_FV_not_appears_indiv.
+ assumption.
+ apply IHl.
+ apply notin_app in H.
+ destruct H as [_ H].
+ assumption.
+
+ simpl.
+ intros.
+ apply notin_app in H.
+ destruct H as [H1 H2].
+ apply Bool.orb_false_intro; auto.
+
+ simpl.
+ intros.
+ apply notin_app in H.
+ destruct H as [H1 H2].
+ apply Bool.orb_false_intro; auto.
+
+ simpl.
+ intros.
+ auto.
+
+ simpl.
+ intros.
+ auto.
+ Defined.
+
+ Lemma fvar_swap_cxt_idem : forall T:Type, forall GD:list (T*formula),
+ forall x y, x\notin FV_cxt GD ->
+ fvar_swap_cxt x GD y = GD.
+ Proof.
+ induction GD.
+
+ reflexivity.
+
+ intros.
+ simpl.
+ rewrite IHGD.
+ rewrite swap_appears_not_typ.
+ destruct a; auto.
+ simpl in H.
+ apply notin_FV_not_appears.
+ apply notin_app in H.
+ apply H.
+ simpl in H.
+ apply notin_app in H.
+ apply H.
+ Defined.
+
+ Lemma fvar_swap_typ_idem : forall C:formula,
+ forall x y, x\notin FV_typ C ->
+ fvar_swap_typ x C y = C.
+ Proof.
+ intros.
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ apply notin_FV_not_appears.
+ assumption.
+ Defined.
+
+ Lemma subst_fvar_swap_lem : forall A,
+ forall x y, x\notin FV_typ A ->
+ fvar_swap_typ x (A^x) y = A^y.
+ Proof.
+ intros.
+ replace (A^y) with ((fvar_swap_typ x A y)^^(fvar_swap_indiv x (fvar x) y)).
+ apply subst_fvar_swap.
+ simpl.
+ destruct var_free_dec; [idtac|congruence].
+ rewrite swap_appears_not_typ.
+ reflexivity.
+ apply notin_FV_not_appears.
+ assumption.
+ Defined.
+
+ Definition fvar_swap_proof : forall (x y:var_free),
+ (forall Gamma t A, proof Gamma t A ->
+ let Gamma' := fvar_swap_cxt x Gamma y in
+ let A' := fvar_swap_typ x A y in
+ proof Gamma' t A').
+ Proof.
+ intros x y.
+ apply proof_rect.
+
+ intros; simpl in *; constructor; auto.
+
+ intros; simpl in *; constructor; auto.
+
+ intros; simpl in *; constructor; auto.
+
+ intros Gamma a A IH.
+ constructor.
+ induction Gamma.
+ inversion IH.
+ simpl in IH.
+ destruct IH.
+ rewrite e.
+ simpl.
+ auto.
+ simpl.
+ auto.
+
+ intros.
+ simpl in *.
+ apply DisjE with (fvar_swap_typ x A y) (fvar_swap_typ x B y);
+ auto.
+
+ intros; simpl in *.
+ apply ImpE with (fvar_swap_typ x A y); auto.
+
+ intros; simpl in *.
+ apply AllI with (x::L).
+ replace (All (fvar_swap_typ x A y)) with (fvar_swap_typ x (All A) y);
+ [idtac | reflexivity].
+ unfold locl.
+ intros.
+ replace k with (plus k 0).
+ apply locl_fvar_swap.
+ intros.
+ apply l.
+ apply Arith.Plus.plus_0_r.
+ intros x0 Hx0.
+ assert (x0 \notin L).
+ clear -Hx0.
+ simpl in *.
+ intro H.
+ auto.
+ replace (fvar_swap_typ x A y ^ x0) with
+ ((fvar_swap_typ x A y) ^^ (fvar_swap_indiv x (fvar x0) y)).
+ rewrite <- subst_fvar_swap.
+ auto.
+ rewrite swap_appears_not_indiv.
+ reflexivity.
+ clear -Hx0.
+ simpl.
+ destruct var_free_dec.
+ simpl in Hx0.
+ assert (Hx0' := Hx0 (inl _ e)).
+ inversion Hx0'.
+ reflexivity.
+
+ intros; simpl in *.
+ rewrite subst_fvar_swap.
+ apply AllE.
+ unfold locli.
+ intros.
+ replace k with (plus k 0).
+ apply locli_fvar_swap.
+ intros.
+ apply l.
+ apply Arith.Plus.plus_0_r.
+ assumption.
+
+ intros; simpl in *.
+ rewrite subst_fvar_swap in H.
+ apply ExI with (fvar_swap_indiv x d y).
+ unfold locli.
+ intros.
+ replace k with (plus k 0).
+ apply locli_fvar_swap.
+ intros.
+ apply l.
+ apply Arith.Plus.plus_0_r.
+ assumption.
+
+ intros; simpl in *.
+ apply ExE with (fvar_swap_typ x A y) (x::L).
+ replace (Ex (fvar_swap_typ x A y)) with (fvar_swap_typ x (Ex A) y);
+ [idtac | reflexivity].
+ unfold locl.
+ intros.
+ replace k with (plus k 0).
+ apply locl_fvar_swap.
+ intros.
+ apply l.
+ apply Arith.Plus.plus_0_r.
+ assumption.
+ intros x0 Hx0.
+ assert (x0 \notin L).
+ clear -Hx0.
+ simpl in *.
+ intro H.
+ auto.
+ replace (fvar_swap_typ x A y ^ x0) with
+ ((fvar_swap_typ x A y) ^^ (fvar_swap_indiv x (fvar x0) y)).
+ rewrite <- subst_fvar_swap.
+ auto.
+ rewrite swap_appears_not_indiv.
+ reflexivity.
+ clear -Hx0.
+ simpl.
+ destruct var_free_dec.
+ simpl in Hx0.
+ assert (Hx0' := Hx0 (inl _ e)).
+ inversion Hx0'.
+ reflexivity.
+ Defined.
+
+ Lemma proof_trm_quant_invar : forall Gamma A,
+ let L := FV_typ A ++ FV_cxt Gamma in
+ (forall x, x\notin L ->
+ sigT (fun t:term => proof Gamma t (A^x))) ->
+ sigT (fun s:term => forall y, y\notin L -> proof Gamma s (A^y)).
+ Proof.
+ intros.
+ destruct (fresh_fvar L) as [x Hx].
+ destruct (H x) as [t p].
+ assumption.
+ exists t.
+ clear H.
+ unfold L in Hx.
+ apply notin_app in Hx.
+ destruct Hx as [Hx1 Hx2].
+ intros y Hy.
+ unfold L in Hy.
+ apply notin_app in Hy.
+ destruct Hy as [Hy1 Hy2].
+ assert (p' := (fvar_swap_proof x y) _ _ _ p).
+ simpl in p'.
+ rewrite fvar_swap_cxt_idem in p'; auto.
+ rewrite subst_fvar_swap_lem in p'; auto.
+ Defined.
+
+ Lemma proof_trm_quant_invar' : forall Gamma a A C,
+ let L := FV_typ C ++ FV_cxt ((a,A)::Gamma) in
+ (forall x, x\notin L ->
+ sigT (fun t:term => proof ((a,A^x)::Gamma) t C)) ->
+ sigT (fun s:term => forall y, y\notin L -> proof ((a,A^y)::Gamma) s C).
+ Proof.
+ intros.
+ destruct (fresh_fvar L) as [x Hx].
+ destruct (H x) as [t p].
+ assumption.
+ exists t.
+ clear H.
+ unfold L in Hx.
+ apply notin_app in Hx.
+ destruct Hx as [Hx1 Hx2].
+ simpl in Hx2.
+ apply notin_app in Hx2.
+ destruct Hx2 as [Hx2 Hx3].
+ intros y Hy.
+ unfold L in Hy.
+ apply notin_app in Hy.
+ destruct Hy as [Hy1 Hy2].
+ simpl in Hy2.
+ apply notin_app in Hy2.
+ destruct Hy2 as [Hy2 Hy3].
+ assert (p' := (fvar_swap_proof x y) _ _ _ p).
+ simpl in p'.
+ rewrite fvar_swap_cxt_idem in p'; auto.
+ rewrite (fvar_swap_typ_idem y Hx1) in p'.
+ rewrite subst_fvar_swap_lem in p'; auto.
+ Defined.
+
+ Lemma proof_subst_lem1 : forall Gamma A, locl (All A) ->
+ let L := FV_typ A ++ FV_cxt Gamma in
+ (forall x, x\notin L -> {t:term & proof Gamma t (A^x)}) ->
+ forall d, locli d -> {t:term & proof Gamma t (A^^d)}.
+ Proof.
+ intros Gamma A Hlocl L H d Hd.
+ apply proof_trm_quant_invar in H.
+ destruct H as [s Hs].
+ apply AllI in Hs.
+ assert (Hs' := AllE Hd Hs).
+ exists (QApp (QLam s)).
+ assumption.
+ assumption.
+ Defined.
+
+ Lemma proof_subst_lem2 : forall Gamma A, locl (Ex A) ->
+ let L := FV_typ A ++ FV_cxt Gamma in
+ {d:indiv & locli d * {t:term & proof Gamma t (A^^d)}} ->
+ {x:var_free & (x\notin L) * {t:term & proof Gamma t (A^x)}}.
+ Proof.
+ intros Gamma A Hlocl L H.
+ destruct H as [d [Hd [t Ht]]].
+ apply ExI in Ht.
+ destruct (fresh_fvar L) as [x Hx].
+ exists x.
+ split.
+ assumption.
+ Admitted.
+End subst_lemmas.
+
+(** The Universal model *)
+Section Universal_Model_and_Completeness.
+ Definition Kworld : Set := context.
+ Definition Kle (w w':Kworld) : Type := incl w w'.
+ Lemma Kle_refl : forall w, Kle w w.
+ Proof.
+ apply incl_refl.
+ Defined.
+
+ Lemma Kle_trans : forall w w' w'', Kle w w' -> Kle w' w'' -> Kle w w''.
+ Proof.
+ unfold Kle.
+ intros.
+ intuition eauto using incl_tran.
+ Defined.
+
+ Definition Normal_term (A:formula)(Gamma:context) :=
+ {t:term & proof Gamma t A}.
+ Hint Unfold Normal_term.
+
+ Definition KX (w:Kworld)(A:formula) : Set := Normal_term A w.
+ Hint Unfold KX.
+
+ Definition Kaforces (w:Kworld)(p:predicate)(ld:list indiv) : Set :=
+ Normal_term (@Atom p ld) w.
+ Hint Unfold Kaforces.
+
+ Notation "w <= w'" := (Kle w w').
+
+ Lemma proof_mon : forall A t w, proof w t A ->
+ forall w', w <= w' -> proof w' t A.
+ Proof.
+ intros A t w p.
+ induction p.
+
+ intros.
+ constructor.
+ auto.
+
+ intros.
+ constructor.
+ auto.
+
+ intros.
+ constructor.
+ apply IHp.
+ apply cons_incl_head.
+ assumption.
+
+ intros.
+ constructor.
+ apply X0.
+ assumption.
+
+ intros.
+ apply DisjE with A B.
+ auto.
+ apply IHp2.
+ apply cons_incl_head.
+ assumption.
+ apply IHp3.
+ apply cons_incl_head.
+ assumption.
+
+ intros.
+ apply ImpE with A.
+ apply IHp1.
+ assumption.
+ apply IHp2.
+ assumption.
+
+ intros.
+ apply AllI with L.
+ assumption.
+ intros.
+ apply X0.
+ assumption.
+ assumption.
+
+ intros.
+ apply AllE.
+ assumption.
+ apply IHp;auto.
+
+ intros.
+ apply ExI with d.
+ assumption.
+ apply IHp.
+ assumption.
+
+ intros.
+ apply ExE with A L.
+ assumption.
+ auto.
+ intros.
+ apply X0 with x.
+ assumption.
+ apply cons_incl_head.
+ assumption.
+ Defined.
+
+ Lemma Kaforces_mon : forall w P ld, @Kaforces w P ld ->
+ forall w', Kle w w' -> @Kaforces w' P ld.
+ Proof.
+ intros w P ld H.
+ destruct H as [v p].
+ intros w' Hle.
+ exists v.
+ unfold Kle in Hle.
+ eauto using proof_mon.
+ Defined.
+
+ Definition K : Kripke.
+ (* begin show *)
+ apply Build_Kripke with Kworld Kle Kaforces.
+ exact Kle_refl.
+ exact Kle_trans.
+ exact KX.
+ exact Kaforces_mon.
+ (* end show *)
+ Defined.
+
+ Let F := (@sforces K).
+
+ Notation "w <= w'" := (Kle w w').
+ Notation "w ||+ A " := (F w A) (at level 70).
+ Notation "w ||- A " := (Kont F w A) (at level 70).
+ Notation " w ||-- Gamma " := (@forces_cxt K w Gamma) (at level 70).
+
+ (** The proof of completeness is via a reify-reflect pair, by induction on [n] i.e. by induction on the complexity of the formula [A]. [fresh] is a fresh variable counter that is increased at every recursive call. *)
+ Theorem Kcompleteness : forall (n:nat)(A:formula), le (depth A) n -> locl A ->
+ forall (w:K)(fresh:nat),
+ (Kont F w A -> {t:term & proof w t A}) *
+ (forall e:term, proof w e A -> Kont F w A).
+ Proof.
+ induction n.
+
+ intros A Hn.
+ destruct A;inversion Hn.
+ intros A HSn Hlocl.
+ intros.
+ destruct A.
+
+ (* Atomic formula *****************)
+ split.
+
+ (* λc. run c *)
+ intros c.
+ apply c.
+ apply wle_refl.
+ simpl.
+ intros w2 w_w2 H.
+ apply H.
+
+ (* λe. ret e *)
+ intros e He.
+ apply ret.
+ simpl.
+ exists e.
+ assumption.
+
+ (* Implication ********************)
+ assert (Hdepth2 : le (depth A2) n).
+ clear -HSn.
+ simpl in HSn.
+ apply le_S_n in HSn.
+ eapply le_trans.
+ apply le_max_r.
+ apply HSn.
+ assert (Hdepth1 : le (depth A1) n).
+ clear -HSn.
+ simpl in HSn.
+ apply le_S_n in HSn.
+ eapply le_trans.
+ apply le_max_l.
+ apply HSn.
+ assert (Hlocl1 : locl A1).
+ clear -Hlocl.
+ unfold locl in *.
+ simpl in *.
+ intros.
+ assert (Hlocl' := Hlocl k d).
+ congruence.
+ assert (Hlocl2 : locl A2).
+ clear -Hlocl.
+ unfold locl in *.
+ simpl in *.
+ intros.
+ assert (Hlocl' := Hlocl k d).
+ congruence.
+ split.
+
+ (* λc.c(λf.λx.↓² f (↑¹_(x,A1),… x) *)
+ intro c.
+ apply c.
+ apply wle_refl.
+ intros w2 le_w2 f.
+ set (w' := (fresh,A1)::w2).
+ assert (f' := sforces_correct_Imp f).
+ clear f.
+ rename f' into f.
+ assert (Hf : Kont F w' A2).
+ apply f.
+ unfold w'.
+ red;simpl;red;auto.
+ apply (snd (IHn _ Hdepth1 Hlocl1 w' (S fresh))) with (Var fresh).
+ unfold w'.
+ constructor.
+ auto.
+ apply (fst (IHn _ Hdepth2 Hlocl2 w' (S fresh))) in Hf.
+ destruct Hf as [t' p'].
+ exists (Lam fresh t').
+ apply ImpI.
+ unfold w' in p'; simpl in p'.
+ assumption.
+
+ (* λe.ret(λa.↑² App e (↓¹ a)) *)
+ intros e He.
+ apply ret.
+ apply sforces_correct_Imp'.
+ intros w' le_w' Ha1.
+ apply (fst (IHn _ Hdepth1 Hlocl1 w' (S fresh))) in Ha1.
+ destruct Ha1 as [a1 Ha1].
+ apply (snd (IHn _ Hdepth2 Hlocl2 w' (S fresh))) with (App e a1).
+ apply ImpE with A1.
+ eauto using proof_mon.
+ assumption.
+
+ (* Disjunction ************************)
+ assert (Hdepth2 : le (depth A2) n).
+ clear -HSn.
+ simpl in HSn.
+ apply le_S_n in HSn.
+ eapply le_trans.
+ apply le_max_r.
+ apply HSn.
+ assert (Hdepth1 : le (depth A1) n).
+ clear -HSn.
+ simpl in HSn.
+ apply le_S_n in HSn.
+ eapply le_trans.
+ apply le_max_l.
+ apply HSn.
+ assert (Hlocl1 : locl A1).
+ clear -Hlocl.
+ unfold locl in *.
+ simpl in *.
+ intros.
+ assert (Hlocl' := Hlocl k d).
+ congruence.
+ assert (Hlocl2 : locl A2).
+ clear -Hlocl.
+ unfold locl in *.
+ simpl in *.
+ intros.
+ assert (Hlocl' := Hlocl k d).
+ congruence.
+ split.
+
+ (* λc.c(λH.match H with inl H₁ => ↓¹ H₁ | inr H₂ => ↓² H₂ *)
+ intro c.
+ apply c.
+ apply wle_refl.
+ intros w2 ww2 Hdisj.
+ unfold F in Hdisj.
+ apply sforces_correct_Disj in Hdisj.
+ destruct Hdisj as [HA1 | HA2].
+ apply (fst (IHn _ Hdepth1 Hlocl1 _ (S fresh))) in HA1.
+ destruct HA1 as [t1 Ht1].
+ exists (Inl t1).
+ constructor.
+ assumption.
+ apply (fst (IHn _ Hdepth2 Hlocl2 _ (S fresh))) in HA2.
+ destruct HA2 as [t2 Ht2].
+ exists (Inr t2).
+ constructor.
+ assumption.
+
+ (* λc.λk. Match e x (k (inl ↑¹_(x,A₁),… x)) x (k (inr ↑²_(x,A2),… x)) *)
+ intros e He.
+ intros C w' le_w' k.
+ simpl in k.
+ assert (k' : forall w2 : Kworld, w' <= w2 -> Kont F w2 A1 + Kont F w2 A2 -> KX w2 C).
+ clear -k.
+ intros w2 w'w2 HDisj.
+ apply sforces_correct_Disj' in HDisj.
+ auto.
+ clear k.
+ rename k' into k.
+ set (w1 := (fresh,A1)::w').
+ assert (branch1 : Kont F w1 A1).
+ apply (snd (IHn _ Hdepth1 Hlocl1 w1 (S fresh))) with (Var fresh).
+ constructor.
+ unfold w1.
+ auto.
+ set (w2 := (fresh,A2)::w').
+ assert (branch2 : Kont F w2 A2).
+ apply (snd (IHn _ Hdepth2 Hlocl2 w2 (S fresh))) with (Var fresh).
+ constructor.
+ unfold w2.
+ auto.
+ assert (w'w1 : w' <= w1).
+ unfold w1.
+ red;simpl;intuition.
+ assert (w'w2 : w' <= w2).
+ unfold w2.
+ red;simpl;intuition.
+ assert (t1p1 := k w1 w'w1 (inl _ branch1)).
+ assert (t2p2 := k w2 w'w2 (inr _ branch2)).
+ destruct t1p1 as [t1 p1].
+ destruct t2p2 as [t2 p2].
+ exists (Match e fresh t1 fresh t2).
+ apply DisjE with A1 A2.
+ apply proof_mon with w.
+ assumption.
+ assumption.
+ assumption.
+ assumption.
+
+ (* Universal quantifier ******************************)
+ split.
+
+ intro c.
+ set (L := FV_typ A ++ FV_cxt w).
+ assert (H0 : forall y:var_free, y\notin L -> KX w (A^y)).
+ intros y Hy.
+ apply c.
+ apply wle_refl.
+ intros w2 ww2 HAll.
+ assert (HAll':= sforces_correct_All HAll).
+ clear HAll.
+ simpl.
+ assert (HAd : Kont F w2 (A^y)).
+ apply HAll'.
+ apply wle_refl.
+ apply locli_fvar.
+ assert (Hdepth' : le (depth (A^y)) n).
+ clear -HSn.
+ simpl in HSn.
+ rewrite depth_subst.
+ apply le_S_n in HSn.
+ assumption.
+ assert (Hlocl' : locl (A^y)).
+ apply locl_locli_subst'.
+ assumption.
+ apply locli_fvar.
+ apply (fst (IHn _ Hdepth' Hlocl' _ (S fresh))) in HAd.
+ destruct HAd as [td Htd].
+ exists td.
+ assumption.
+ apply proof_trm_quant_invar in H0.
+ destruct H0 as [s Hs].
+ exists (QLam s).
+ apply AllI with L.
+ assumption.
+ unfold L.
+ assumption.
+
+ intros e He.
+ apply ret.
+ apply sforces_correct_All'.
+ intros w1 ww1 d Hd.
+ assert (Hdepth' : le (depth (A^^d)) n).
+ clear -HSn.
+ simpl in HSn.
+ rewrite depth_subst.
+ apply le_S_n in HSn.
+ assumption.
+ assert (Hlocl' : locl (A^^d)).
+ apply locl_locli_subst'.
+ assumption.
+ assumption.
+ apply (snd (IHn _ Hdepth' Hlocl' _ (S fresh))) with (QApp e).
+ apply AllE.
+ assumption.
+ apply proof_mon with w.
+ assumption.
+ assumption.
+
+ (* Existential quantifier ****************************)
+ split.
+
+ intro c.
+ apply c.
+ apply wle_refl.
+ intros w2 ww2 HEx.
+ unfold F in HEx.
+ apply sforces_correct_Ex in HEx.
+ destruct HEx as [d [Hd HAd]].
+ assert (Hdepth' : le (depth (A^^d)) n).
+ clear -HSn.
+ simpl in HSn.
+ rewrite depth_subst.
+ apply le_S_n in HSn.
+ assumption.
+ assert (Hlocl' : locl (A^^d)).
+ apply locl_locli_subst''.
+ assumption.
+ assumption.
+ apply (fst (IHn _ Hdepth' Hlocl' _ (S fresh))) in HAd.
+ destruct HAd as [t Ht].
+ exists (QWit t).
+ apply ExI with d.
+ assumption.
+ assumption.
+
+ intros e He.
+ intros C w' le_w' k.
+ simpl in k.
+ assert (k' : forall w2 : Kworld, w' <= w2 ->
+ sigT (fun (d:indiv) => locli d * Kont (K:=K) (sforces (K:=K)) w2 (A ^^ d))
+ -> KX w2 C).
+ clear -k.
+ intros w2 w'w2 HEx.
+ apply sforces_correct_Ex' in HEx.
+ auto.
+ clear k.
+ rename k' into k.
+ set (w1 := fun d => fun (w:@world K) => ((fresh, A^^d)::w):Kworld).
+ assert (ww1 : forall d w0, wle w0 (w1 d w0)).
+ intro d.
+ unfold w1.
+ right.
+ assumption.
+ assert (Hdepth' : forall d0, le (depth (A^^d0)) n).
+ clear -HSn.
+ intro d0.
+ simpl in HSn.
+ rewrite depth_subst.
+ apply le_S_n in HSn.
+ assumption.
+ assert (Hlocl' : forall d0, locli d0 -> locl (A^^d0)).
+ apply locl_locli_subst''.
+ assumption.
+ assert (branch1 : forall d, locli d -> Kont F (w1 d w) (A^^d)).
+ intros d Hd.
+ apply (snd (IHn _ (Hdepth' d) (Hlocl' d Hd) (w1 d w) (S fresh))) with (Var fresh).
+ constructor.
+ unfold w1.
+ left; reflexivity.
+ assert (H0 : {s:term & forall y : var_free,
+ y \notin FV_typ C ++ FV_cxt ((fresh,A)::w') -> proof ((fresh,A^y)::w') s C}).
+ apply proof_trm_quant_invar'.
+ intros x Hx.
+ assert ({t : term & proof (w1 (fvar x) w') t C}).
+ apply k.
+ apply ww1.
+ exists (fvar x).
+ split.
+ apply locli_fvar.
+ apply (snd (IHn _ (Hdepth' (fvar x)) (Hlocl' (fvar x) (locli_fvar x)) (w1 (fvar x) w') (S fresh))) with (Var fresh).
+ constructor.
+ unfold w1.
+ left;reflexivity.
+ unfold w1 in H.
+ assumption.
+ destruct H0 as [s Hs].
+ exists (QMatch e fresh s).
+ apply ExE with A (FV_typ C ++ FV_cxt ((fresh,A)::w')).
+ assumption.
+ apply proof_mon with w.
+ assumption.
+ assumption.
+ clear -Hs.
+ intros x Hx.
+ auto.
+ Defined.
+End Universal_Model_and_Completeness.
diff --git a/kripke_completeness/int_comp_noquant.v b/kripke_completeness/int_comp_noquant.v
new file mode 100644
index 0000000..2ffa2fc
--- /dev/null
+++ b/kripke_completeness/int_comp_noquant.v
@@ -0,0 +1,531 @@
+(** Formalised proof of completeness of Intuitionistic propositional
+ logic (disjunction and implication, no quantifiers) with respect to
+ Kripke-style models.
+
+ Danko Ilik, July 2009
+
+ Normalisation examples at end of file.
+*)
+Require Import stdlib_Type.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+(** printing <= $\le$ #≤# *)
+(** printing ||-- $\Vdash$ #⊩# *)
+(** printing ||- $\Vdash$ #⊩# *)
+(** printing ||+ $\Vdash_s$ #⊩s # *)
+
+Definition var : Set := nat.
+
+(** Proof terms *)
+Inductive term : Set :=
+| Lam : var -> term -> term
+| Inl : term -> term
+| Inr : term -> term
+| Var : var -> term
+| App : term -> term -> term
+| Match : term -> var -> term -> var -> term -> term.
+
+Inductive formula :=
+ Atom : formula
+| Imp : formula -> formula -> formula
+| Disj : formula -> formula -> formula.
+
+Definition context := list (var*formula).
+Hint Unfold context.
+
+(** Natural deduction system *)
+Inductive proof : context -> term -> formula -> Set :=
+| DisjIL : forall Gamma t A B, proof Gamma t A -> proof Gamma (Inl t) (Disj A B)
+| DisjIR : forall Gamma t A B, proof Gamma t B -> proof Gamma (Inr t) (Disj A B)
+| ImpI : forall Gamma x t A B, proof ((x,A)::Gamma) t B -> proof Gamma (Lam x t) (Imp A B)
+| Hypo : forall Gamma x A, In (x,A) Gamma -> proof Gamma (Var x) A
+| DisjE : forall Gamma x y e t u A B C, proof Gamma e (Disj A B) -> proof ((x,A)::Gamma) t C -> proof ((y,B)::Gamma) u C -> proof Gamma (Match e x t y u) C
+| ImpE : forall Gamma t u A B, proof Gamma t (Imp A B) ->
+ proof Gamma u A -> proof Gamma (App t u) B.
+
+Section Abstract_Semantics.
+ (** An abstract Kripke-style structure: [wle] is the preorder, [X] is
+ the exploding-world predicate, [aforces] is strong forcing of
+ atomic formulae. *)
+ Record Kripke : Type := {
+ world :> Set;
+ wle : world -> world -> Type;
+ wle_refl : forall w, wle w w;
+ wle_trans : forall w w' w'', wle w w' -> wle w' w'' -> wle w w'';
+ X : world -> formula -> Set;
+ aforces : world -> Set;
+ aforces_mon : forall w, aforces w -> forall w', wle w w' -> aforces w'
+ }.
+ Notation "w <= w'" := (wle w w').
+
+ Variable K:Kripke.
+
+ (** The continuations monad we will use to extend forcing to
+ composite formulas *)
+ Definition Kont (F:K->formula->Type)(w:K)(A:formula) :=
+ forall C:formula,
+ (forall w1, w <= w1 ->
+ (forall w2, w1 <= w2 -> F w2 A -> X w2 C) -> X w1 C).
+ Hint Unfold Kont.
+
+ (** Strong forcing extended to composite formulas. The expression
+ [sforces w A] stands for strong forcing, while [Kont sforces w A]
+ stands for (weak) forcing of A in the world w. *)
+ Fixpoint sforces (w:K)(A:formula) {struct A} : Type :=
+ match A with
+ | Atom => aforces w
+ | Imp A1 A2 => forall w', w <= w' -> Kont sforces w' A1 -> Kont sforces w' A2
+ | Disj A1 A2 => sum (Kont sforces w A1) (Kont sforces w A2)
+ end.
+
+ Lemma sforces_mon : forall A w, sforces w A -> forall w', w <= w' -> sforces w' A .
+ Proof.
+ induction A.
+ simpl.
+ intros.
+ eauto using aforces_mon.
+
+ intros w H1 w' ww'.
+ simpl in *.
+ eauto using wle_trans.
+
+ intros w H1 w' ww'.
+ simpl in *.
+ case H1; intro H'.
+ eauto using wle_trans.
+ eauto using wle_trans.
+ Defined.
+
+ Definition ret {w A} : sforces w A -> Kont sforces w A.
+ Proof.
+ intros H.
+ intros C w1 w_w1 k.
+ apply k.
+ apply wle_refl.
+ apply sforces_mon with w.
+ assumption.
+ assumption.
+ Defined.
+
+ Fixpoint Kont_sforces_cxt (w:K)(Gamma:context) : Type :=
+ match Gamma with
+ | nil => unit
+ | cons aA Gamma' =>
+ prod (Kont sforces w (snd aA)) (Kont_sforces_cxt w Gamma')
+ end.
+
+ Notation " w ||- A " := (Kont sforces w A) (at level 70).
+ Notation " w ||-- Gamma " := (Kont_sforces_cxt w Gamma) (at level 70).
+
+ Lemma Kont_sforces_cxt_mon : forall Gamma w, w ||-- Gamma -> forall w', w <= w' -> w' ||-- Gamma.
+ Proof.
+ induction Gamma.
+ simpl.
+ auto.
+ simpl.
+ intros w H w' w_w'.
+ destruct H.
+ split; eauto using wle_trans,sforces_mon.
+ Defined.
+
+ Lemma Kont_sforces_mon : forall w A, w ||- A -> forall w', w <= w' -> w' ||- A.
+ Proof.
+ intros w A H w' ww'.
+ intros C w'' w'w'' k.
+ apply H.
+ eauto using wle_trans.
+ assumption.
+ Defined.
+
+ Theorem soundness : forall Gamma t A, proof Gamma t A -> forall w, w ||-- Gamma -> w ||- A.
+ Proof.
+ intros Gamma t A p.
+ induction p.
+
+ intros w wGamma.
+ apply ret.
+ simpl.
+ left; auto.
+
+ intros w wGamma.
+ apply ret.
+ simpl.
+ right; auto.
+
+ intros w wGamma.
+ apply ret.
+ simpl in *.
+ intros.
+ apply IHp.
+ split; eauto using Kont_sforces_cxt_mon.
+
+ induction Gamma.
+ simpl.
+ contradict i.
+ simpl.
+ intros w [Ha HGamma].
+ simpl in i.
+ destruct i.
+ rewrite e in Ha.
+ assumption.
+ auto.
+
+ clear p1 p2 p3.
+ intros w wGamma.
+ intros D w1 ww1 k.
+ apply IHp1 with w.
+ assumption.
+ assumption.
+ intros w2 w1w2 HDisj.
+ simpl in HDisj.
+ case HDisj.
+ intro HA.
+ apply IHp2 with w2.
+ simpl.
+ split.
+ assumption.
+ eauto using wle_trans,Kont_sforces_cxt_mon.
+ apply wle_refl.
+ eauto using wle_trans.
+ intro HB.
+ apply IHp3 with w2.
+ simpl.
+ split.
+ assumption.
+ eauto using wle_trans,Kont_sforces_cxt_mon.
+ apply wle_refl.
+ eauto using wle_trans.
+
+ intros w wGamma.
+ intros C w1 ww1 k.
+ apply IHp1 with w1.
+ eauto using Kont_sforces_cxt_mon.
+ apply wle_refl.
+ simpl sforces.
+ intros w2 w1w2 HAB.
+ apply HAB with w2.
+ apply wle_refl.
+ eauto using wle_trans,Kont_sforces_mon.
+ apply wle_refl.
+ intros w3 w2w3 HB.
+ eauto using wle_trans,Kont_sforces_mon.
+ Defined.
+End Abstract_Semantics.
+
+(** The Universal model *)
+Section Universal_Model_and_Completeness.
+ Definition Kworld : Set := context.
+ Definition Kle (w w':Kworld) : Type := incl w w'.
+ Lemma Kle_refl : forall w, Kle w w.
+ Proof.
+ apply incl_refl.
+ Defined.
+
+ Lemma Kle_trans : forall w w' w'', Kle w w' -> Kle w' w'' -> Kle w w''.
+ Proof.
+ unfold Kle.
+ intros.
+ intuition eauto using incl_tran.
+ Defined.
+
+ Definition Normal_term (A:formula)(Gamma:context) :=
+ {t:term & proof Gamma t A}.
+ Hint Unfold Normal_term.
+
+ Definition KX (w:Kworld)(A:formula) : Set := Normal_term A w.
+ Hint Unfold KX.
+
+ Definition Kaforces (w:Kworld) : Set := Normal_term Atom w.
+
+ Notation "w <= w'" := (Kle w w').
+
+ Lemma proof_mon : forall A t w, proof w t A -> forall w', w <= w' -> proof w' t A.
+ Proof.
+ intros A t w p.
+ induction p.
+
+ intros.
+ constructor.
+ auto.
+
+ intros.
+ constructor.
+ auto.
+
+ intros.
+ constructor.
+ apply IHp.
+ apply cons_incl_head.
+ assumption.
+
+ intros.
+ constructor.
+ apply X0.
+ assumption.
+
+ intros.
+ apply DisjE with A B.
+ auto.
+ apply IHp2.
+ apply cons_incl_head.
+ assumption.
+ apply IHp3.
+ apply cons_incl_head.
+ assumption.
+
+ intros.
+ apply ImpE with A.
+ apply IHp1.
+ assumption.
+ apply IHp2.
+ assumption.
+ Defined.
+
+ Lemma Kaforces_mon : forall w, Kaforces w -> forall w', Kle w w' -> Kaforces w'.
+ Proof.
+ intros w H.
+ destruct H as [v p].
+ intros w' Hle.
+ exists v.
+ unfold Kle in Hle.
+ eauto using proof_mon.
+ Defined.
+
+ Definition K : Kripke.
+ (* begin show *)
+ apply Build_Kripke with Kworld Kle Kaforces.
+ exact Kle_refl.
+ exact Kle_trans.
+ exact KX.
+ exact Kaforces_mon.
+ (* end show *)
+ Defined.
+
+ Let F := (@sforces K).
+
+ (** The proof of completeness is via a reify-reflect pair. [fresh] is a fresh variable counter that is increased at every recursive call. *)
+ Theorem Kcompleteness : forall A w, forall fresh:nat,
+ (Kont F w A -> {t:term & proof w t A}) *
+ (forall e:term, proof w e A -> Kont F w A).
+ Proof.
+ induction A.
+
+ (* Atomic formula *)
+ intros.
+ split.
+
+ intros c.
+ apply c.
+ apply wle_refl.
+ simpl.
+ intros w2 w_w2 H.
+ apply H.
+
+ intros e He.
+ apply ret.
+ simpl.
+ exists e.
+ assumption.
+
+ (* Implication *)
+ intro w.
+ split.
+
+ intro c.
+ apply c.
+ apply wle_refl.
+ intros w2 le_w2 f.
+ set (w' := (fresh,A1)::w2).
+ simpl in f.
+ assert (Hf : Kont F w' A2).
+ apply f.
+ unfold w'.
+ red;simpl;auto.
+ apply (snd (IHA1 w' (S fresh))) with (Var fresh).
+ unfold w'.
+ constructor.
+ auto.
+ apply (fst (IHA2 w' (S fresh))) in Hf.
+ destruct Hf as [t' p'].
+ exists (Lam fresh t').
+ apply ImpI.
+ unfold w' in p'; simpl in p'.
+ assumption.
+
+ intros e He.
+ apply ret.
+ simpl.
+ intros w' le_w' Ha1.
+ apply (fst (IHA1 w' (S fresh))) in Ha1.
+ destruct Ha1 as [a1 Ha1].
+ apply (snd (IHA2 w' (S fresh))) with (App e a1).
+ apply ImpE with A1.
+ eauto using proof_mon.
+ assumption.
+
+ (* Disjunction *)
+ intro w.
+ split.
+
+ intro c.
+ apply c.
+ apply wle_refl.
+ intros w2 ww2 Hdisj.
+ simpl in Hdisj.
+ destruct Hdisj as [HA1 | HA2].
+ apply (fst (IHA1 _ (S fresh))) in HA1.
+ destruct HA1 as [t1 Ht1].
+ exists (Inl t1).
+ constructor.
+ assumption.
+ apply (fst (IHA2 _ (S fresh))) in HA2.
+ destruct HA2 as [t2 Ht2].
+ exists (Inr t2).
+ constructor.
+ assumption.
+
+ intros e He.
+ intros C w' le_w' k.
+ simpl in k.
+ set (w1 := (fresh,A1)::w').
+ assert (branch1 : Kont F w1 A1).
+ apply (snd (IHA1 w1 (S fresh))) with (Var fresh).
+ constructor.
+ unfold w1.
+ auto.
+ set (w2 := (fresh,A2)::w').
+ assert (branch2 : Kont F w2 A2).
+ apply (snd (IHA2 w2 (S fresh))) with (Var fresh).
+ constructor.
+ unfold w2.
+ auto.
+ assert (w'w1 : w' <= w1).
+ unfold w1.
+ red;simpl;intuition.
+ assert (w'w2 : w' <= w2).
+ unfold w2.
+ red;simpl;intuition.
+ assert (t1p1 := k w1 w'w1 (inl _ branch1)).
+ assert (t2p2 := k w2 w'w2 (inr _ branch2)).
+ destruct t1p1 as [t1 p1].
+ destruct t2p2 as [t2 p2].
+ exists (Match e fresh t1 fresh t2).
+ apply DisjE with A1 A2.
+ apply proof_mon with w.
+ assumption.
+ assumption.
+ assumption.
+ assumption.
+ Defined.
+End Universal_Model_and_Completeness.
+
+(** A lemma by which NbE below uses Kcompeteness *)
+Lemma Kont_sforces_cxt_refl : forall Gamma, Kont_sforces_cxt (K:=K) Gamma Gamma.
+Proof.
+ induction Gamma.
+ simpl.
+ constructor.
+ simpl.
+ split.
+ destruct a as [x A].
+ simpl.
+ apply (snd (Kcompleteness _ _ 0)) with (Var x).
+ constructor.
+ left;reflexivity.
+ apply Kont_sforces_cxt_mon with Gamma.
+ assumption.
+ simpl.
+ right.
+ assumption.
+Defined.
+
+Definition NbE (Gamma:context)(t:term)(A:formula)(p:proof Gamma t A) : term.
+Proof.
+(* begin show *)
+ intros.
+ set (compl := fst (Kcompleteness A Gamma 0)).
+ set (sndns := soundness (K:=K) p (Kont_sforces_cxt_refl Gamma)).
+ exact (projT1 (compl sndns)).
+(* end show *)
+Defined.
+
+(** Some tests of the normalization algorithm *)
+
+Definition id1 : proof nil (Lam 0 (Var 0)) (Imp Atom Atom).
+Proof.
+ apply ImpI.
+ apply Hypo.
+ left; reflexivity.
+Defined.
+
+(* Eval compute in (NbE id1). *)
+
+Definition id2 : proof nil (Lam 0 (Var 0)) (Imp (Imp Atom Atom) (Imp Atom Atom)).
+Proof.
+ apply ImpI.
+ apply Hypo.
+ left; reflexivity.
+Defined.
+
+(* Eval compute in (NbE id2). *)
+
+Definition id3 : proof nil (Lam 0 (Var 0)) (Imp (Imp (Imp Atom Atom) (Imp Atom Atom)) (Imp (Imp Atom Atom) (Imp Atom Atom))).
+Proof.
+ apply ImpI.
+ apply Hypo.
+ left; reflexivity.
+Defined.
+
+(* Eval compute in (NbE id3). *)
+
+Definition id4 : proof nil (Lam 0 (Var 0)) (Imp (Disj Atom Atom) (Disj Atom Atom)).
+Proof.
+ apply ImpI.
+ apply Hypo.
+ left; reflexivity.
+Defined.
+
+(* Eval compute in (NbE id4). *)
+
+Definition id5 : proof nil (Lam 0 (Var 0)) (Imp (Imp (Disj Atom Atom) (Disj Atom Atom)) (Imp (Disj Atom Atom) (Disj Atom Atom))).
+Proof.
+ apply ImpI.
+ apply Hypo.
+ left; reflexivity.
+Defined.
+
+(* Eval compute in (NbE id5). *)
+
+Definition id6 : proof nil (Lam 0 (Var 0)) (Imp (Imp (Disj Atom Atom) Atom) (Imp (Disj Atom Atom) Atom)).
+Proof.
+ apply ImpI.
+ apply Hypo.
+ left; reflexivity.
+Defined.
+
+(* Eval compute in (NbE id6). *)
+
+Definition id7 : proof nil (Lam 0 (Lam 1 (Match (Var 1) 2 (App (Var 0) (Inl (Var 2))) 2 (App (Var 0) (Inr (Var 2)))))) (Imp (Imp (Disj Atom Atom) Atom) (Imp (Disj Atom Atom) Atom)).
+Proof.
+ apply ImpI.
+ apply ImpI.
+ apply DisjE with Atom Atom.
+ apply Hypo.
+ left; reflexivity.
+ apply ImpE with (Disj Atom Atom).
+ apply Hypo.
+ right; right; left; reflexivity.
+ constructor.
+ apply Hypo.
+ left; reflexivity.
+ apply ImpE with (Disj Atom Atom).
+ apply Hypo.
+ right; right; left; reflexivity.
+ constructor.
+ apply Hypo.
+ left; reflexivity.
+Defined.
+
+(* Eval compute in (NbE id7). *)
diff --git a/kripke_completeness/stdlib_Type.v b/kripke_completeness/stdlib_Type.v
new file mode 100644
index 0000000..dd36e07
--- /dev/null
+++ b/kripke_completeness/stdlib_Type.v
@@ -0,0 +1,196 @@
+Require Export Max.
+Require Export List.
+
+Set Implicit Arguments.
+
+(** The [le] and [max] of Coq Standard Library live in Prop, while we need it in Type. *)
+
+Section le_max_Type.
+
+ Inductive le (n:nat) : nat -> Set :=
+ | le_n : le n n
+ | le_S : forall m:nat, le n m -> le n (S m)
+ where "n <= m" := (le n m) : nat_scope.
+
+ Lemma le_trans : forall n m p, le n m -> le m p -> le n p.
+ Proof.
+ induction 2; auto.
+ constructor.
+ assumption.
+ Defined.
+ Hint Resolve le_trans: arith v62.
+
+ Theorem le_n_S : forall n m, n <= m -> S n <= S m.
+ Proof.
+ induction 1; auto.
+ constructor.
+ constructor.
+ assumption.
+ Defined.
+
+ Theorem le_S_n : forall n m, S n <= S m -> n <= m.
+ Proof.
+ intros n m H; change (pred (S n) <= pred (S m)) in |- *.
+ destruct H; simpl; auto with arith.
+ constructor.
+ apply le_trans with (S n).
+ constructor.
+ constructor.
+ assumption.
+ Defined.
+
+ Theorem le_O_n : forall n, 0 <= n.
+ Proof.
+ induction n.
+ constructor.
+ constructor.
+ assumption.
+ Defined.
+
+ Lemma max_l : forall n m, m <= n -> max n m = n.
+ Proof.
+ induction n; induction m; simpl in |- *; auto with arith.
+ intro H.
+ inversion H.
+ intro H.
+ assert (max n m = n).
+ apply IHn.
+ inversion H.
+ constructor.
+ eapply le_trans.
+ apply le_S.
+ apply le_n.
+ assumption.
+ rewrite H0.
+ reflexivity.
+ Defined.
+
+ Lemma max_r : forall n m, n <= m -> max n m = m.
+ Proof.
+ induction n; induction m; simpl in |- *; auto with arith.
+ intro H.
+ inversion H.
+ intro H.
+ assert (max n m = m).
+ apply IHn.
+ inversion H.
+ constructor.
+ eapply le_trans.
+ apply le_S.
+ apply le_n.
+ assumption.
+ rewrite H0.
+ reflexivity.
+ Defined.
+
+ Lemma le_max_l : forall n m, n <= max n m.
+ Proof.
+ double induction n m; intros; auto with arith.
+ constructor.
+ constructor.
+ assumption.
+ constructor.
+ simpl.
+ apply le_n_S.
+ auto.
+ Defined.
+
+ Lemma le_max_r : forall n m, m <= max n m.
+ Proof.
+ induction n; simpl in |- *; auto with arith.
+ intro.
+ constructor.
+ induction m; simpl in |- *; auto with arith.
+ apply le_O_n.
+ apply le_n_S.
+ auto.
+ Defined.
+End le_max_Type.
+
+(** The [In] relation of [List] from Coq Standard Library lives in
+ Prop, while we need it and some related lemmas, in Type. *)
+Section In_Type.
+
+ Variable A:Type.
+
+ Open Scope type_scope.
+
+ Fixpoint In (a:A) (l:list A) {struct l} : Set :=
+ match l with
+ | nil => Empty_set
+ | b :: m => (b = a) + In a m
+ end.
+
+ Definition incl (l m:list A) := forall a:A, In a l -> In a m.
+ Hint Unfold incl.
+
+ Lemma incl_refl : forall l:list A, incl l l.
+ Proof.
+ auto.
+ Defined.
+ Hint Resolve incl_refl.
+
+ Lemma incl_tran : forall l m n:list A, incl l m -> incl m n -> incl l n.
+ Proof.
+ auto.
+ Defined.
+
+ Hint Unfold incl.
+ Hint Resolve incl_refl.
+ Lemma cons_incl_head : forall l1 l2:list A, forall a:A,
+ incl l1 l2 -> incl (a::l1) (a::l2).
+ Proof.
+ intros l1 l2 a H.
+ red;simpl.
+ intros b H1.
+ destruct H1.
+ left; assumption.
+ right.
+ apply H.
+ assumption.
+ Defined.
+
+ Lemma in_or_app : forall (l m:list A) (a:A), In a l + In a m -> In a (l ++ m).
+ Proof.
+ intros l m a.
+ elim l; simpl in |- *; intro H.
+ now_show (In a m).
+ elim H; auto; intro H0.
+ now_show (In a m).
+ elim H0. (* subProof completed *)
+ intros y H0 H1.
+ now_show ((H = a) + In a (y ++ m)).
+ elim H1; auto 4.
+ intro H2.
+ now_show ((H = a) + In a (y ++ m)).
+ elim H2; auto.
+ Defined.
+
+ Variable f : A -> bool.
+
+ Lemma forallb_forall :
+ forall l, forallb f l = true <-> (forall x, In x l -> f x = true).
+ Proof.
+ induction l; simpl; intuition.
+ destruct (andb_prop _ _ H1).
+ congruence.
+ destruct (andb_prop _ _ H1); auto.
+ assert (forallb f l = true).
+ apply H0; intuition.
+ rewrite H1; auto.
+ Defined.
+
+End In_Type.
+Hint Unfold incl.
+Hint Unfold In.
+Hint Resolve incl_refl.
+
+(** The [iff] relation of Coq Standard Library lives in
+ Prop, while we need it in Type. *)
+Section iff_Type.
+ Open Scope type_scope.
+ Definition iff (A B:Type) := (A -> B) * (B -> A).
+End iff_Type.
+Notation "A <-> B" := (iff A B) : type_scope.
+Hint Unfold iff: extcore.
+
diff --git a/ssrFOL01.v b/ssrFOL01.v
index f3a8537..f43fd9b 100644
--- a/ssrFOL01.v
+++ b/ssrFOL01.v
@@ -39,6 +39,17 @@ Notation "x → y" := (implS x y) (at level 51).
Notation "∀ x , y" := (forallS x y) (at level 51).
Notation "∃ x , y" := (existsS x y) (at level 51).
+Definition allFml (p : Fml -> bool) f :=
+ match f with
+ | notS f1 => p f1
+ | orS f1 f2 => p f1 && p f2
+ | andS f1 f2 => p f1 && p f2
+ | implS f1 f2 => p f1 && p f2
+ | forallS x f1 => p f1
+ | existsS x f1 => p f1
+ | _ => true
+ end.
+
(* Syntactic comparison and eqType for Fml. *)
@@ -59,8 +70,8 @@ Lemma eqFmlP : Equality.axiom eqFml.
Proof.
move=> F G.
apply (iffP idP); last first; [move => -> //| ].
- - elim: G => //.
- + move=> v v0 /=.
+ - elim: G => //; move=>*; by apply/andP.
+(* + move=> v v0 /=.
by apply/andP.
+ move=> v v0 /=.
by apply/andP.
@@ -73,10 +84,15 @@ Proof.
+ move=> v f H => /=.
by apply/andP.
+ move=> v f H => /=.
- by apply/andP.
-
- - move: G; elim F.
- + move=> v v0; case => v1 v2 //=.
+ by apply/andP.*)
+
+ - move: G; elim F; (move=> f H f0 H0 G E || move=> v f H G E || move => v v0 G E); move: G E;
+ solve [ case=> v1 v2 //= /andP [] /eqP -> /eqP -> //
+ | case=> f // /v0 -> //
+ | case=> f1 f2 //= /andP [] /H -> /H0 -> //
+ | case=> v0 f0 //= /andP [] /eqP -> /H -> //].
+ (*case=> ? ? //= /andP; solve [elim=> /eqP -> /eqP -> // | elim=> /H -> /H0 -> // | elim=> /eqP -> /H -> //]*)
+(* + move=> v v0; case => v1 v2 //=.
by move/andP; elim => /eqP -> /eqP ->.
+ move=> v v0; case => v1 v2 //=.
by move/andP; elim => /eqP -> /eqP ->.
@@ -91,9 +107,29 @@ Proof.
+ move=> v f H; case => v0 f0 //=.
by move/andP; elim => /eqP -> /H ->.
+ move=> v f H; case => v0 f0 //=.
- by move/andP; elim => /eqP -> /H ->.
+ by move/andP; elim => /eqP -> /H ->.*)
Qed.
+(* See: *)
+(* Handbook of Practical Logic, 6.6 Proving tautologies by inference *)
+(* coq-8.4pl5/tactics/tauto.ml4 *)
+
+(* propositional tautology *)
+Fixpoint tautoFml F G {struct F} :=
+ match F, G with
+ | equalityS m0 m1, equalityS n0 n1 => (m0 == n0) && (m1 == n1)
+ | membershipS m0 m1, membershipS n0 n1 => (m0 == n0) && (m1 == n1)
+ | notS F0, notS G0 => eqFml F0 G0
+ | orS F0 F1, orS G0 G1 => eqFml F0 G0 && eqFml F1 G1
+ | andS F0 F1, andS G0 G1 => eqFml F0 G0 && eqFml F1 G1
+ | implS F0 F1, implS G0 G1 => eqFml F0 G0 && eqFml F1 G1
+ | forallS m F0, forallS n G0 => (m == n) && eqFml F0 G0
+ | existsS m F0, existsS n G0 => (m == n) && eqFml F0 G0
+ | _, _ => false
+ end.
+
+
+
Canonical Fml_eqMixin := EqMixin eqFmlP.
Canonical Fml_eqType := Eval hnf in EqType Fml Fml_eqMixin.
@@ -159,6 +195,13 @@ Definition is_not_qf (f : Fml) : bool :=
| _ => true
end.
+Fixpoint is_quantifier_free (f : Fml) : bool :=
+ match f with
+ | forallS _ _ => false
+ | existsS _ _ => false
+ | _ => allFml is_quantifier_free f
+ end.
+
Fixpoint is_quantifier_free (f : Fml) : bool :=
match f with
| equalityS x y => true