Library ModelHF


A finitary model of the calculus of construction: hereditarily finite sets.

Require Import Bool List.
Require Import HFcoc.
Require Import Models GenModelSyntax.

Instance of the abstract model


Module HF_Coc_Model <: CC_Model.

Definition X := hf.
Definition inX := In_hf.
Definition eqX := Eq_hf.
Definition eqX_equiv : Equivalence Eq_hf.
Qed.

Notation "x ∈ y" := (inX x y).
Notation "x == y" := (eqX x y).

Lemma in_ext: Proper (eqX ==> eqX ==> iff) inX.
Definition props := props.
Definition app := cc_app.
Definition lam := cc_lam.
Definition prod := cc_prod.

Definition eq_fun (x:X) (f1 f2:X->X) :=
   y1 y2, y1 x -> y1 == y2 -> f1 y1 == f2 y2.

Lemma lam_ext :
   x1 x2 f1 f2,
  x1 == x2 ->
  eq_fun x1 f1 f2 ->
  lam x1 f1 == lam x2 f2.
Lemma app_ext: Proper (eqX ==> eqX ==> eqX) app.
Lemma prod_ext :
   x1 x2 f1 f2,
  x1 == x2 ->
  eq_fun x1 f1 f2 ->
  prod x1 f1 == prod x2 f2.
Lemma prod_intro : dom f F,
  eq_fun dom f f ->
  eq_fun dom F F ->
  ( x, x dom -> f x F x) ->
  lam dom f prod dom F.
Lemma prod_elim : dom f x F,
  eq_fun dom F F ->
  f prod dom F ->
  x dom ->
  app f x F x.
Lemma impredicative_prod : dom F,
  eq_fun dom F F ->
  ( x, x dom -> F x props) ->
  prod dom F props.
Lemma beta_eq:
   dom F x,
  eq_fun dom F F ->
  x dom ->
  app (lam dom F) x == F x.
End HF_Coc_Model.

Module Soundness := GenModelSyntax.MakeModel(HF_Coc_Model).

Import Soundness.
Import T.

Principles supported by the model

Deciding whether f is true for all valuations
Fixpoint forall_int_env (e:Env.env) (f:(nat->hf)->bool) {struct e} : bool :=
  match e with
    nil => f (fun _ => empty)
  | T::e' => forall_int_env e'
      (fun i_f => forall_elt (fun x => f (V.cons x i_f)) (int (int_trm T) i_f))
  end.

Deciding whether a context is satisfiable
Definition valid_context (e:Env.env) : bool :=
  negb (forall_int_env e (fun _ => false)).


Require Import Term TypeJudge.

Definition int_clos T := int (int_trm T) vnil.

Consistency of CC

  Lemma cc_consistency : M M', ~ eq_typ nil M M' FALSE.

Properties not validated by the model

Eval vm_compute in (int_clos FALSE).

Properties validated by the model

Eval vm_compute in (int_clos TRUE).

NAT is inhabited
Eval vm_compute in (int_clos INAT).
excluded-middle:
Eval vm_compute in (int_clos EM).
proof-irrelevance:
Eval vm_compute in (int_clos PI).
extensionality of funs of Prop -> Prop
Eval vm_compute in (int_clos (EXT prop prop)).