added papers
This commit is contained in:
parent
eac2ecb81e
commit
030e7f0a6d
64 changed files with 37112 additions and 0 deletions
BIN
papers/altenkirch-reuss-monadic.pdf
Normal file
BIN
papers/altenkirch-reuss-monadic.pdf
Normal file
Binary file not shown.
BIN
papers/barendregt-convention.pdf
Normal file
BIN
papers/barendregt-convention.pdf
Normal file
Binary file not shown.
10928
papers/barendregt-pure.pdf
Normal file
10928
papers/barendregt-pure.pdf
Normal file
File diff suppressed because it is too large
Load diff
BIN
papers/chapman-thesis.pdf
Normal file
BIN
papers/chapman-thesis.pdf
Normal file
Binary file not shown.
BIN
papers/conor-outrageous.pdf
Normal file
BIN
papers/conor-outrageous.pdf
Normal file
Binary file not shown.
BIN
papers/conor-turing.pdf
Normal file
BIN
papers/conor-turing.pdf
Normal file
Binary file not shown.
BIN
papers/conor-universe.pdf
Normal file
BIN
papers/conor-universe.pdf
Normal file
Binary file not shown.
254
papers/domains-2009/KnasterTarski.v
Executable file
254
papers/domains-2009/KnasterTarski.v
Executable file
|
@ -0,0 +1,254 @@
|
|||
Require Import PredomCore.
|
||||
Require Import PredomProd.
|
||||
|
||||
Section KnasterTarski.
|
||||
|
||||
Section CLattice.
|
||||
|
||||
(** printing <= $\sqsubseteq$ *)
|
||||
(** printing sup $\bigwedge$ *)
|
||||
(** printing inf $\bigvee$ *)
|
||||
(** printing == $\equiv$ *)
|
||||
|
||||
(** * Complete Lattice
|
||||
A semi-lattice [(L, \/ )] is a partial order [L = (T, <= )] and a binary operation inf [\/ : L -> L -> L],
|
||||
such that [forall x y, x \/ y <= x], [forall x y, x \/ y == y \/ x] and [forall z x y, z <= x /\ z <= y -> z <= x \/ y].
|
||||
|
||||
A complete lattice is a semi lattice where the inf operation has been generalized to all subsets of [L].
|
||||
*)
|
||||
|
||||
Record CLattice := mk_semilattice
|
||||
{
|
||||
ltord :> ord;
|
||||
inf : forall S : ltord -> Prop, ltord;
|
||||
Pinf : forall P : ltord -> Prop, forall t, P t -> inf P <= t;
|
||||
PinfL : forall P : ltord -> Prop, forall t, (forall x, P x -> t <= x) -> t <= inf P
|
||||
}.
|
||||
|
||||
Variable L : CLattice.
|
||||
|
||||
(** A complete lattice has both infima and suprema. Suprema of a set S ([sup S]) is given by [inf { z | S <= z}] *)
|
||||
|
||||
Definition sup (S:L -> Prop) : L := inf L (fun D => forall s, S s -> s <= D).
|
||||
|
||||
Lemma Psup : forall (S:L -> Prop) t, S t -> t <= sup S.
|
||||
intros S t St. unfold sup.
|
||||
refine (PinfL _ _ _ _).
|
||||
intros l Cl. apply Cl. apply St.
|
||||
Qed.
|
||||
|
||||
Lemma PsupL : forall (S:L -> Prop) t, (forall x, S x -> x <= t) -> sup S <= t.
|
||||
intros S t C. unfold sup. refine (Pinf _ _ _ _).
|
||||
intros l Sl. apply C. apply Sl.
|
||||
Qed.
|
||||
|
||||
End CLattice.
|
||||
|
||||
Section SubLattice.
|
||||
Variable L : CLattice.
|
||||
|
||||
(** Let [P] represent a subset of [L]. *)
|
||||
Variable P : ltord L -> Prop.
|
||||
|
||||
(** We can then define a an ordering on the subset [P] by inheriting the ordering from [L]. *)
|
||||
Definition Subord : ord.
|
||||
exists ({t | P t}) (fun (x y:{t | P t}) => match x with exist x _ =>
|
||||
match y with exist y _ => x <= y
|
||||
end
|
||||
end).
|
||||
intros x. case x. clear x. intros x _. auto.
|
||||
intros x y z ; case x ; case y ; case z ; clear x z y. intros x _ y _ z _.
|
||||
apply Ole_trans.
|
||||
Defined.
|
||||
|
||||
Definition ForgetSubord : Subord -> ltord L.
|
||||
intros S. case S. clear S. intros x _. apply x.
|
||||
Defined.
|
||||
|
||||
(** If [P] is closed under infima then it is a complete lattice *)
|
||||
Variable PSinf: forall (S:ltord L -> Prop), (forall t, S t -> P t) -> P (inf L S).
|
||||
|
||||
Definition Embed (t:ltord L) (p:P t) : Subord.
|
||||
intros t Pt. exists t. apply Pt.
|
||||
Defined.
|
||||
|
||||
Definition Extend (S : Subord -> Prop) : ltord L -> Prop.
|
||||
intros S l. apply (exists p:P l, S (Embed l p)).
|
||||
Defined.
|
||||
|
||||
Lemma ExtendP: forall (S:Subord -> Prop) t, Extend S t -> P t.
|
||||
intros S t Ex. unfold Extend in Ex. unfold Embed in Ex. destruct Ex as [E _].
|
||||
apply E.
|
||||
Defined.
|
||||
|
||||
Definition SubLattice : CLattice.
|
||||
exists (Subord) (fun (S:Subord -> Prop) =>
|
||||
Embed (inf L (Extend S)) (PSinf (Extend S) (ExtendP S))).
|
||||
intros S t. case t. simpl. clear t. intros l Pl Pe. refine (Pinf _ _ _ _ ).
|
||||
unfold Extend. unfold Embed. exists Pl. apply Pe.
|
||||
intros S t. case t. simpl. clear t. intros x Px C. apply PinfL.
|
||||
intros t et.
|
||||
unfold Extend in et. unfold Embed in et.
|
||||
destruct et as [Pt St]. specialize (C _ St). apply C.
|
||||
Defined.
|
||||
|
||||
End SubLattice.
|
||||
|
||||
Variable L : CLattice.
|
||||
|
||||
Definition op_ord (O:ord) : ord.
|
||||
intros O. exists (tord O) (fun x y => y <= x).
|
||||
auto. intros x y z. intros A B. apply (Ole_trans B A).
|
||||
Defined.
|
||||
|
||||
Definition op_lat : CLattice.
|
||||
exists (op_ord L) (sup L).
|
||||
intros P t Pt. simpl. apply (Psup L P _ Pt).
|
||||
intros P t C. simpl. refine (PsupL L P _ _). apply C.
|
||||
Defined.
|
||||
|
||||
Variable f : L -m> L.
|
||||
|
||||
Definition PreFixedPoint d := f d <= d.
|
||||
|
||||
Lemma PreFixedInfs: (forall S : L -> Prop,
|
||||
(forall t : L, S t -> PreFixedPoint t) -> PreFixedPoint (inf L S)).
|
||||
intros S C.
|
||||
unfold PreFixedPoint in *.
|
||||
apply (PinfL L S). intros l Sl.
|
||||
apply Ole_trans with (y:= f l).
|
||||
assert (monotonic f) as M by auto. apply M.
|
||||
apply Pinf. apply Sl.
|
||||
apply (C _ Sl).
|
||||
Qed.
|
||||
|
||||
Definition PreFixedLattice := SubLattice L (fun x => PreFixedPoint x) PreFixedInfs.
|
||||
|
||||
Definition PostFixedPoint d := d <= f d.
|
||||
|
||||
Lemma PostFixedSups: forall S : op_lat -> Prop,
|
||||
(forall t : op_lat, S t -> (fun x : op_lat => PostFixedPoint x) t) ->
|
||||
(fun x : op_lat => PostFixedPoint x) (inf op_lat S).
|
||||
unfold PostFixedPoint.
|
||||
intros S C. simpl in *. refine (PsupL L _ _ _).
|
||||
intros l Sl. apply Ole_trans with (y:= f l). apply C. apply Sl.
|
||||
assert (monotonic f) as M by auto. apply M. apply (Psup L _ _ Sl).
|
||||
Qed.
|
||||
|
||||
Definition PostFixedLattice := SubLattice op_lat (fun x => PostFixedPoint x) PostFixedSups.
|
||||
|
||||
Definition lfp := match (inf PreFixedLattice (fun _ => True)) with exist x _ => x end.
|
||||
|
||||
Definition gfp := match (inf PostFixedLattice (fun _ => True)) with exist x _ => x end.
|
||||
|
||||
Lemma lfp_fixedpoint : f lfp == lfp.
|
||||
unfold lfp.
|
||||
assert (yy:= Pinf PreFixedLattice (fun _ => True)).
|
||||
generalize yy ; clear yy.
|
||||
case (inf PreFixedLattice (fun _ => True)).
|
||||
intros x Px yy.
|
||||
assert (zz:=fun X PX => (yy (exist (fun t => PreFixedPoint t) X PX))). clear yy.
|
||||
assert (PreFixedPoint (f x)). unfold PreFixedPoint.
|
||||
assert (monotonic f) as M by auto. apply M. apply Px.
|
||||
split.
|
||||
specialize (zz _ H). simpl in zz.
|
||||
apply Px.
|
||||
apply zz ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma gfp_fixedpoint : f gfp == gfp.
|
||||
unfold gfp.
|
||||
assert (yy:= Pinf PostFixedLattice (fun _ => True)).
|
||||
generalize yy ; clear yy.
|
||||
case (inf PostFixedLattice (fun _ => True)).
|
||||
intros x Px yy.
|
||||
assert (zz:=fun X PX => (yy (exist (fun t => PostFixedPoint t) X PX))). clear yy.
|
||||
assert (PostFixedPoint (f x)). unfold PostFixedPoint.
|
||||
assert (monotonic f) as M by auto. apply M. apply Px.
|
||||
split.
|
||||
specialize (zz _ H). simpl in zz. apply zz. auto.
|
||||
apply Px.
|
||||
Qed.
|
||||
|
||||
Lemma lfp_least: forall fp, f fp <= fp -> lfp <= fp.
|
||||
unfold lfp.
|
||||
intros fp ff. assert (PreFixedPoint fp). apply ff.
|
||||
assert (yy:= Pinf PreFixedLattice (fun _ => True)).
|
||||
generalize yy ; clear yy.
|
||||
case (inf PreFixedLattice (fun _ => True)).
|
||||
intros x Px yy.
|
||||
assert (zz:=fun X PX => (yy (exist (fun t => PreFixedPoint t) X PX))). clear yy.
|
||||
case (inf PreFixedLattice (fun _ : PreFixedLattice => True)).
|
||||
specialize (zz _ H). simpl in zz. intros _ _. apply zz. auto.
|
||||
Qed.
|
||||
|
||||
Lemma gfp_greatest: forall fp, fp <= f fp -> fp <= gfp.
|
||||
unfold gfp.
|
||||
intros fp ff. assert (PostFixedPoint fp). apply ff.
|
||||
assert (yy:= Pinf PostFixedLattice (fun _ => True)).
|
||||
generalize yy ; clear yy.
|
||||
case (inf PostFixedLattice (fun _ => True)).
|
||||
intros x Px yy.
|
||||
assert (zz:=fun X PX => (yy (exist (fun t => PostFixedPoint t) X PX))). clear yy.
|
||||
case (inf PostFixedLattice (fun _ : PostFixedLattice => True)).
|
||||
specialize (zz _ H). simpl in zz. intros _ _. apply zz. auto.
|
||||
Qed.
|
||||
|
||||
End KnasterTarski.
|
||||
|
||||
Definition ProdLattice (L1 L2 : CLattice) : CLattice.
|
||||
intros L1 L2. exists (Oprod L1 L2) (fun (S:Oprod L1 L2 -> Prop) =>
|
||||
(inf L1 (fun l1 => exists l2, S (l1,l2)),inf L2 (fun l2 => exists l1, S (l1,l2)))).
|
||||
intros P t. case t. clear t. intros t0 t1 Pt. simpl. split.
|
||||
refine (Pinf _ _ _ _). exists t1. apply Pt.
|
||||
refine (Pinf _ _ _ _). exists t0. apply Pt.
|
||||
intros P t. case t. clear t. intros t0 t1 Pt. simpl. split.
|
||||
refine (PinfL _ _ _ _). intros x Px. destruct Px as [y Pxy].
|
||||
specialize (Pt _ Pxy).
|
||||
inversion Pt. simpl in *. auto.
|
||||
|
||||
refine (PinfL _ _ _ _). intros x Px. destruct Px as [y Pyx].
|
||||
specialize (Pt _ Pyx). inversion Pt. simpl in *. auto.
|
||||
Defined.
|
||||
|
||||
Definition LFst L1 L2 : ProdLattice L1 L2 -m> L1.
|
||||
intros L1 L2. exists (@fst L1 L2). unfold monotonic.
|
||||
intros x y ; case x ; clear x ; intros x0 x1 ; case y ; clear y ; intros y0 y1.
|
||||
simpl. intros [A _] ; auto.
|
||||
Defined.
|
||||
|
||||
Definition LSnd L1 L2 : ProdLattice L1 L2 -m> L2.
|
||||
intros L1 L2. exists (@snd L1 L2). unfold monotonic.
|
||||
intros x y ; case x ; clear x ; intros x0 x1 ; case y ; clear y ; intros y0 y1.
|
||||
simpl. intros [_ A] ; auto.
|
||||
Defined.
|
||||
|
||||
Definition ArrowLattice (T:Type) (L:CLattice) : CLattice.
|
||||
intros T L. exists (ford T L) (fun (S:T -o> L -> Prop) => fun t => inf L (fun st => exists s, S s /\ s t = st)).
|
||||
intros P f Pf. simpl. intros t. refine (Pinf _ _ _ _). exists f. split. auto. auto.
|
||||
intros P f C. simpl. intros t. refine (PinfL _ _ _ _).
|
||||
intros l Tl. destruct Tl as [f0 [Pf0 f0t]]. specialize (C _ Pf0). rewrite <- f0t.
|
||||
auto.
|
||||
Defined.
|
||||
|
||||
Require Import Sets.
|
||||
|
||||
Definition Subsetord (T:SetKind) : ord.
|
||||
intros T. exists (set T) (@subset T).
|
||||
intros x t ; auto.
|
||||
intros x y z C1 C2 t. specialize (C1 t). specialize (C2 t).
|
||||
intros a. apply (C2 (C1 a)).
|
||||
Defined.
|
||||
|
||||
Definition SubsetLattice (T:SetKind) : CLattice.
|
||||
intros T. exists (Subsetord T) (@Intersect T).
|
||||
intros P t Pt. simpl. intros x Px. apply (Px t Pt).
|
||||
intros P t Pt. simpl in *. intros x Px f Pf. specialize (Pt f Pf).
|
||||
apply Pt. apply Px.
|
||||
Defined.
|
||||
|
||||
Definition Supsetord (T:SetKind) := op_ord (Subsetord T).
|
||||
|
||||
Definition SupsetLattice (T:SetKind) := op_lat (SubsetLattice T).
|
||||
|
||||
|
50
papers/domains-2009/Makefile
Executable file
50
papers/domains-2009/Makefile
Executable file
|
@ -0,0 +1,50 @@
|
|||
COQFLAGS=-coqlib "$(COQLIB)"
|
||||
COQDOCFLAGS=-coqlib_path "$(COQLIB)" --latex -g
|
||||
COQC="$(COQBIN)"/coqc $(COQFLAGS)
|
||||
COQDOC="$(COQBIN)"/coqdoc $(COQDOCFLAGS)
|
||||
|
||||
|
||||
.SUFFIXES: .v .vo .tex
|
||||
|
||||
.v.vo:
|
||||
$(COQC) $*
|
||||
.v.tex:
|
||||
$(COQDOC) $*.v
|
||||
|
||||
|
||||
typedlambda.tex: typedlambda.v
|
||||
typedopsem.tex: typedopsem.v
|
||||
typeddensem.tex: typeddensem.v
|
||||
typedsubst.tex: typedsubst.v
|
||||
typedsoundness.tex: typedsoundness.v
|
||||
typedadequacy.tex: typedadequacy.v
|
||||
|
||||
utility.vo: utility.v
|
||||
Sets.vo: Sets.v
|
||||
|
||||
PredomCore.vo: PredomCore.v
|
||||
PredomProd.vo: PredomProd.v PredomCore.vo
|
||||
PredomSum.vo: PredomSum.v PredomCore.vo
|
||||
PredomLift.vo: PredomLift.v PredomCore.vo utility.vo
|
||||
PredomKleisli.vo: PredomKleisli.v PredomLift.vo PredomProd.vo PredomCore.vo utility.vo
|
||||
PredomFix.vo: PredomFix.v PredomCore.vo PredomProd.vo PredomLift.vo PredomKleisli.vo
|
||||
KnasterTarski.vo: KnasterTarski.v Sets.vo PredomCore.vo PredomProd.vo
|
||||
lc.vo: lc.v PredomCore.vo PredomProd.vo PredomSum.vo PredomLift.vo PredomFix.vo PredomKleisli.vo
|
||||
PredomAll.vo: PredomAll.v PredomCore.vo PredomProd.vo PredomSum.vo PredomLift.vo PredomKleisli.vo PredomFix.vo
|
||||
|
||||
uni.vo: uni.v utility.vo
|
||||
unisem.vo: unisem.v PredomAll.vo lc.vo uni.vo
|
||||
uniade.vo: uniade.v unisem.vo KnasterTarski.vo Sets.vo PredomAll.vo
|
||||
unisound.vo: unisound.v unisem.vo PredomAll.vo
|
||||
|
||||
typedlambda.vo: typedlambda.v utility.vo
|
||||
typedopsem.vo: typedopsem.v typedlambda.vo
|
||||
typeddensem.vo: typeddensem.v PredomAll.vo typedlambda.vo
|
||||
typedsubst.vo: typedsubst.v typeddensem.vo
|
||||
typedsoundness.vo: typedsoundness.v typedsubst.vo typedopsem.vo
|
||||
typedadequacy.vo: typedadequacy.v typeddensem.vo typedopsem.vo
|
||||
|
||||
all: typedadequacy.vo typedsoundness.vo unisound.vo uniade.vo
|
||||
|
||||
|
||||
|
7
papers/domains-2009/PredomAll.v
Executable file
7
papers/domains-2009/PredomAll.v
Executable file
|
@ -0,0 +1,7 @@
|
|||
Require Export PredomCore.
|
||||
Require Export PredomProd.
|
||||
Require Export PredomSum.
|
||||
Require Export PredomLift.
|
||||
Require Export PredomKleisli.
|
||||
Require Export PredomFix.
|
||||
Implicit Arguments eta [D].
|
1430
papers/domains-2009/PredomCore.v
Executable file
1430
papers/domains-2009/PredomCore.v
Executable file
File diff suppressed because it is too large
Load diff
388
papers/domains-2009/PredomFix.v
Executable file
388
papers/domains-2009/PredomFix.v
Executable file
|
@ -0,0 +1,388 @@
|
|||
(*==========================================================================
|
||||
Definition of fixpoints and associated lemmas
|
||||
==========================================================================*)
|
||||
Require Import PredomCore.
|
||||
Require Import PredomProd.
|
||||
Require Import PredomLift.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
(** ** Fixpoints *)
|
||||
|
||||
Section Fixpoints.
|
||||
Variable D : cpo.
|
||||
Context {dp : Pointed D}.
|
||||
|
||||
(* Variable D: cpo. *)
|
||||
Variable f : D -m>D.
|
||||
|
||||
Hypothesis fcont : continuous f.
|
||||
|
||||
Fixpoint iter_ n (* D [Pointed D] (f : D -m> D) *) : D := match n with O => PBot | S m => f (iter_ m (* f *)) end.
|
||||
|
||||
Lemma iter_incr (* D [Pointed D] (f : D -m> D)*) : forall n, iter_ n (* f *) <= f (iter_ n (*f *)).
|
||||
induction n; simpl; auto.
|
||||
(* apply Pleast. *)
|
||||
Save.
|
||||
|
||||
Hint Resolve iter_incr.
|
||||
|
||||
Definition iter : natO -m> D.
|
||||
exists iter_; red; intros.
|
||||
induction H; simpl; auto.
|
||||
apply Ole_trans with (iter_ m); auto.
|
||||
Defined.
|
||||
|
||||
Definition fixp : D := lub iter.
|
||||
|
||||
Lemma fixp_le : fixp <= f fixp.
|
||||
unfold fixp.
|
||||
apply Ole_trans with (lub (fmon_comp f iter)); auto.
|
||||
Save.
|
||||
Hint Resolve fixp_le.
|
||||
|
||||
Lemma fixp_eq : fixp == f fixp.
|
||||
apply Ole_antisym; auto.
|
||||
unfold fixp.
|
||||
apply Ole_trans with (1:= fcont iter).
|
||||
rewrite (lub_lift_left iter (S O)); auto.
|
||||
Save.
|
||||
|
||||
Lemma fixp_inv : forall g, f g <= g -> fixp <= g.
|
||||
unfold fixp; intros.
|
||||
apply lub_le.
|
||||
induction n; intros; auto.
|
||||
simpl; apply Ole_trans with (f g); auto.
|
||||
Save.
|
||||
|
||||
End Fixpoints.
|
||||
Hint Resolve fixp_le fixp_eq fixp_inv.
|
||||
|
||||
|
||||
Definition fixp_cte D {pd:Pointed D} : forall (d:D), fixp (fmon_cte D d) == d.
|
||||
intros; apply fixp_eq with (f:=fmon_cte D d); red; intros; auto.
|
||||
unfold fmon_cte at 1; simpl.
|
||||
apply (le_lub (fmon_comp (fmon_cte D (O2:=D) d) c) (O)); simpl; auto.
|
||||
Save.
|
||||
Hint Resolve fixp_cte.
|
||||
|
||||
|
||||
Lemma fixp_le_compat D {pd:Pointed D} : forall (*(D:cpo)*) (f g : D-m>D), f<=g -> fixp f <= fixp g.
|
||||
intros; unfold fixp.
|
||||
apply lub_le_compat.
|
||||
intro n; induction n; simpl; auto.
|
||||
apply Ole_trans with (g (iter_ (D:=D) f n)); auto.
|
||||
Save.
|
||||
Hint Resolve fixp_le_compat.
|
||||
|
||||
|
||||
Add Parametric Morphism D {pd : Pointed D} : (@fixp D pd)
|
||||
with signature (@Oeq (D -m>D)) ==> (@Oeq D)
|
||||
as fixp_eq_compat.
|
||||
intros; apply Ole_antisym; auto.
|
||||
Save.
|
||||
Hint Resolve fixp_eq_compat.
|
||||
|
||||
Definition Fixp D {pd: Pointed D} : (D-m>D) -m> D.
|
||||
intros; exists (@fixp D pd (* (D:=D) *)).
|
||||
auto.
|
||||
Defined.
|
||||
|
||||
Lemma Fixp_simpl D {pd:Pointed D} : forall (* (D:cpo)*) (f:D-m>D), Fixp f = fixp f.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Definition Iter D {pd:Pointed D} : (D-M->D) -m> (natO -M->D).
|
||||
intros; exists (fun (f:D-M->D) => iter (D:=D) f); red; intros.
|
||||
intro n; induction n; simpl; intros; auto.
|
||||
apply Ole_trans with (y (iter_ (D:=D) x n)); auto.
|
||||
Defined.
|
||||
|
||||
Lemma IterS_simpl D {pd:Pointed D} : forall (* (D:cpo)*) f n, Iter f (S n) = f (Iter f n).
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma iterS_simpl D {pd:Pointed D} : forall (* (D:cpo)*) f n, iter f (S n) = f (iter f n).
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma iter_continuous D {pd:Pointed D} :(* forall (D:cpo), *)
|
||||
forall h : natO -m> (D-M->D), (forall n, continuous (h n)) ->
|
||||
iter (lub h) <= lub (Iter (* D*) @ h).
|
||||
red; intros; intro k.
|
||||
induction k; auto.
|
||||
(* botbot above just so auto works here *)
|
||||
|
||||
rewrite iterS_simpl.
|
||||
apply Ole_trans with (lub h (lub (c:=natO -M-> D) (Iter (*D*) @ h) k)); auto.
|
||||
repeat (rewrite fmon_lub_simpl).
|
||||
apply Ole_trans with
|
||||
(lub (c:=D) (lub (c:=natO-M->D) (double_app h ((Iter (* D*) @ h) <_> k)))).
|
||||
apply lub_le_compat; simpl; intros.
|
||||
apply Ole_trans with (lub ((h x)@((Iter (*D*) @ h) <_> k))); auto.
|
||||
apply Ole_trans
|
||||
with (lub (fmon_diag (double_app h ((Iter (*D*) @ h) <_> k)))); auto.
|
||||
case (double_lub_diag (double_app h ((Iter (*D*) @ h) <_> k))); intros L _.
|
||||
apply Ole_trans with (1:=L); auto.
|
||||
Save.
|
||||
|
||||
Hint Resolve iter_continuous.
|
||||
|
||||
Lemma iter_continuous_eq D {pd:Pointed D} :
|
||||
forall h : natO -m> (D-M->D), (forall n, continuous (h n)) ->
|
||||
iter (lub h) == lub (Iter (* D*) @ h).
|
||||
intros; apply Ole_antisym; auto.
|
||||
exact (lub_comp_le (Iter (*D*)) h).
|
||||
Save.
|
||||
|
||||
|
||||
Lemma fixp_continuous D {pd:Pointed D} : forall (* (D:cpo)*) (h : natO -m> (D-M->D)),
|
||||
(forall n, continuous (h n)) -> fixp (lub h) <= lub (Fixp (*D*) @ h).
|
||||
intros; unfold fixp.
|
||||
rewrite (iter_continuous_eq H).
|
||||
simpl; rewrite <- lub_exch_eq; auto.
|
||||
Save.
|
||||
Hint Resolve fixp_continuous.
|
||||
|
||||
Lemma fixp_continuous_eq D {pd:Pointed D} : forall (h : natO -m> (D -M-> D)),
|
||||
(forall n, continuous (h n)) -> fixp (lub h) == lub (Fixp (*D*) @ h).
|
||||
intros; apply Ole_antisym; auto.
|
||||
exact (lub_comp_le (D1:=D -M-> D) (Fixp (*D*)) h).
|
||||
Save.
|
||||
|
||||
Definition FIXP (D:cpo) {p:Pointed D} : (D-C->D) -c> D.
|
||||
intros D p.
|
||||
exists (Fixp @ (Fcontit D D)).
|
||||
unfold continuous.
|
||||
intros h.
|
||||
rewrite fmon_comp_simpl.
|
||||
rewrite Fixp_simpl.
|
||||
apply Ole_trans with (fixp (D:=D) (lub (c:=D-M->D) (Fcontit D D@h))); auto.
|
||||
apply Ole_trans with (lub (Fixp (*D*) @ (Fcontit D D@h))); auto.
|
||||
apply fixp_continuous; intros n.
|
||||
change (continuous (D1:=D) (D2:=D) (fcontit (h n))); auto.
|
||||
Defined.
|
||||
|
||||
Lemma FIXP_simpl D {pd:Pointed D} : forall (*(D:cpo)*) (f:D-c>D), FIXP (*D*) f = Fixp (*D*) (fcontit f).
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma FIXP_le_compat D {pd:Pointed D} : forall (*(D:cpo)*) (f g : D-C->D),
|
||||
f <= g -> FIXP (*D*) f <= FIXP (*D*) g.
|
||||
intros; repeat (rewrite FIXP_simpl); repeat (rewrite Fixp_simpl).
|
||||
apply fixp_le_compat; auto.
|
||||
Save.
|
||||
Hint Resolve FIXP_le_compat.
|
||||
|
||||
Add Parametric Morphism (D:cpo) (p:Pointed D) : (@FIXP D p)
|
||||
with signature (@Ole (D -c> D)) ++> (@Ole D)
|
||||
as FIXP_leq_compat.
|
||||
intros; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D:cpo) (p:Pointed D) : (@FIXP D p)
|
||||
with signature (@Oeq (D -c> D)) ==> (@Oeq D)
|
||||
as FIXP_eq_compat.
|
||||
intros; apply Ole_antisym; auto.
|
||||
Qed.
|
||||
|
||||
Hint Resolve FIXP_eq_compat.
|
||||
Hint Resolve FIXP_eq_compat_Morphism.
|
||||
|
||||
Lemma FIXP_eq D {pd:Pointed D} : forall (f:D-c>D), FIXP (*D*) f == f (FIXP (*D*) f).
|
||||
intros; rewrite FIXP_simpl; rewrite Fixp_simpl.
|
||||
apply (fixp_eq (f:=fcontit f)).
|
||||
auto.
|
||||
Save.
|
||||
Hint Resolve FIXP_eq.
|
||||
|
||||
Lemma FIXP_com: forall E D (f:E -C-> (D -C-> D)) {pd:Pointed D}, FIXP << f == EV << <| f, FIXP << f |>.
|
||||
intros E D f P. apply fcont_eq_intro. intros e. repeat (rewrite fcont_comp_simpl).
|
||||
rewrite FIXP_eq. auto.
|
||||
Qed.
|
||||
|
||||
Lemma FIXP_inv D {pd:Pointed D} : forall (f:D-c>D)(g : D), f g <= g -> FIXP (*D*) f <= g.
|
||||
intros; rewrite FIXP_simpl; rewrite Fixp_simpl; apply fixp_inv; auto.
|
||||
Save.
|
||||
|
||||
(** *** Iteration of functional *)
|
||||
Lemma FIXP_comp_com D {pd:Pointed D} : forall (f g:D-c>D),
|
||||
g << f <= f << g-> FIXP (*D*) g <= f (FIXP (*D*) g).
|
||||
intros; apply FIXP_inv.
|
||||
apply Ole_trans with (f (g (FIXP (*D*) g))).
|
||||
assert (X:=fcont_le_elim H (FIXP (*D*) g)). repeat (rewrite fcont_comp_simpl in X). apply X.
|
||||
apply fcont_monotonic.
|
||||
case (FIXP_eq g); trivial.
|
||||
Save.
|
||||
|
||||
Lemma FIXP_comp D {pd:Pointed D} : forall (f g:D-c>D),
|
||||
g << f <= f << g -> f (FIXP (*D*) g) <= FIXP (*D*) g -> FIXP (*D*) (f << g) == FIXP (*D*) g.
|
||||
intros; apply Ole_antisym.
|
||||
(* fix f << g <= fix g *)
|
||||
apply FIXP_inv.
|
||||
rewrite fcont_comp_simpl.
|
||||
apply Ole_trans with (f (FIXP (*D*) g)); auto.
|
||||
(* fix g <= fix f << g *)
|
||||
apply FIXP_inv.
|
||||
assert (g (f (FIXP (*D*) (f << g))) <= f (g (FIXP (*D*) (f << g)))).
|
||||
specialize (H (FIXP (f << g))). repeat (rewrite fcont_comp_simpl in H). apply H.
|
||||
case (FIXP_eq (f<<g)); intros.
|
||||
apply Ole_trans with (2:=H3).
|
||||
apply Ole_trans with (2:=H1).
|
||||
apply fcont_monotonic.
|
||||
apply FIXP_inv.
|
||||
rewrite fcont_comp_simpl.
|
||||
apply fcont_monotonic.
|
||||
apply Ole_trans with (1:=H1); auto.
|
||||
Save.
|
||||
|
||||
Fixpoint fcont_compn D {pd:Pointed D} (f:D-c>D) (n:nat) {struct n} : D-c>D :=
|
||||
match n with O => f | S p => fcont_compn f p << f end.
|
||||
|
||||
Lemma fcont_compn_com D {pd:Pointed D} : forall (f:D-c>D) (n:nat),
|
||||
f << (fcont_compn f n) <= fcont_compn f n << f.
|
||||
induction n; auto.
|
||||
simpl fcont_compn.
|
||||
apply Ole_trans with ((f << fcont_compn (D:=D) f n) << f); auto.
|
||||
Save.
|
||||
|
||||
Lemma FIXP_compn D {pd:Pointed D} :
|
||||
forall (f:D-c>D) (n:nat), FIXP (fcont_compn f n) == FIXP f.
|
||||
induction n; auto.
|
||||
simpl fcont_compn.
|
||||
apply FIXP_comp.
|
||||
apply fcont_compn_com.
|
||||
apply Ole_trans with (fcont_compn (D:=D) f n (FIXP (fcont_compn (D:=D) f n))); auto.
|
||||
apply Ole_trans with (FIXP (fcont_compn (D:=D) f n)); auto.
|
||||
Save.
|
||||
|
||||
Lemma fixp_double D {pd:Pointed D} : forall (f:D-c>D), FIXP (f << f) == FIXP f.
|
||||
intros; exact (FIXP_compn f (S O)).
|
||||
Save.
|
||||
|
||||
|
||||
(** *** Induction principle *)
|
||||
Definition admissible (D:cpo) (P:D->Prop) :=
|
||||
forall f : natO -m> D, (forall n, P (f n)) -> P (lub f).
|
||||
|
||||
Definition fixp_ind D { pd : Pointed D} : forall (F: D -m> D)(P : D-> Prop),
|
||||
admissible P -> P PBot -> (forall x, P x -> P (F x)) -> P (fixp F).
|
||||
intros D pD F P Adm Pb Pc; unfold fixp.
|
||||
apply Adm. intros n.
|
||||
induction n; simpl; auto.
|
||||
Defined.
|
||||
|
||||
Definition SubOrd (D:ord) (P:D -> Prop) : ord.
|
||||
intros D I. exists ({x : tord D | I x})
|
||||
(fun (x' y:{x : D | I x}) => match x' with exist x _ => match y with exist y
|
||||
_ => @Ole D x y end end).
|
||||
intros x1. case x1. auto.
|
||||
intros x1. case x1. clear x1. intros x ix y1. case y1. clear y1.
|
||||
intros y iy z1. case z1. clear z1. intros z iz l1 l2.
|
||||
apply (Ole_trans l1 l2).
|
||||
Defined.
|
||||
|
||||
Definition SubOrde (D:cpo) (P:D -> Prop) (d:D) : P d -> SubOrd P.
|
||||
intros. exists d. auto.
|
||||
Defined.
|
||||
|
||||
Definition InheritFunm (D E:cpo) (Q:E->Prop) (f:D -m> E) (p:forall d, Q (f d)) :
|
||||
D -m> SubOrd Q.
|
||||
intros D E Q f p.
|
||||
exists (fun d => @SubOrde E Q (f d) (p d)).
|
||||
unfold monotonic. intros x y lxy. auto.
|
||||
Defined.
|
||||
|
||||
Definition Forgetm D P : (@SubOrd D P) -m> D.
|
||||
intros D P. exists (fun (O:SubOrd P) => match O with exist x _ => x end).
|
||||
unfold monotonic. intros x. case x. clear x. intros x px y. case y. clear y.
|
||||
intros ; auto.
|
||||
Defined.
|
||||
|
||||
Definition Subchainlub (D:cpo) (P:D -> Prop) (I:admissible P) (c:natO -m> (SubOrd
|
||||
P)) : SubOrd P.
|
||||
intros D P I c.
|
||||
exists (@lub D (Forgetm P @ c)).
|
||||
unfold admissible in I. specialize (I (Forgetm P @ c)).
|
||||
apply I. intros i. simpl. case (c i). auto.
|
||||
Defined.
|
||||
|
||||
Definition SubCPO (D:cpo) (P:D -> Prop) (I: admissible P) : cpo.
|
||||
intros D P I. exists (SubOrd P) (Subchainlub I).
|
||||
intros c n. unfold Subchainlub. simpl. case_eq (c n). intros x Px cn.
|
||||
refine (Ole_trans _ (@le_lub _ (Forgetm P @ c) n)).
|
||||
simpl. rewrite cn. auto.
|
||||
intros c d C. unfold Subchainlub. simpl.
|
||||
case_eq d. intros dd pd de.
|
||||
refine (lub_le _). intros n. simpl. specialize (C n).
|
||||
case_eq (c n). intros d1 pd1 cn. rewrite cn in C.
|
||||
simpl in C. rewrite de in C. auto.
|
||||
Defined.
|
||||
|
||||
Definition InheritFun (D E:cpo) (Q:E -> Prop) (IQ : admissible Q) (f:D -c> E)
|
||||
(p:forall d, Q (f d)) : D -c> SubCPO IQ.
|
||||
intros. exists (InheritFunm p). unfold continuous.
|
||||
intros c. simpl.
|
||||
rewrite lub_comp_eq.
|
||||
apply lub_le. intros n.
|
||||
refine (Ole_trans _ (le_lub (Forgetm Q @ (InheritFunm p @ c)) n)).
|
||||
auto. auto.
|
||||
Defined.
|
||||
|
||||
Lemma InheritFun_simpl: forall (D E:cpo) (Q:E -> Prop) (IQ : admissible Q) (f:D -c> E)
|
||||
(p:forall d, Q (f d)) d, InheritFun IQ p d = InheritFunm p d.
|
||||
intros. auto.
|
||||
Qed.
|
||||
|
||||
Definition Forget (D:cpo) (P:D -> Prop) (I: admissible P) : (SubCPO I) -c> D.
|
||||
intros D P I. exists (@Forgetm D P).
|
||||
unfold continuous. auto.
|
||||
Defined.
|
||||
|
||||
Lemma ForgetP D P fs d : P (@Forget D P fs d).
|
||||
intros. case d. auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D:cpo) (P:D -> Prop) (I:admissible P) : (@Forget D P I)
|
||||
with signature (@Ole _) ++> (@Ole _)
|
||||
as Forget_le_compat.
|
||||
intros d0 d1 dle. auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D:cpo) (P:D -> Prop) (I:admissible P) : (@Forget D P I)
|
||||
with signature (@Oeq _) ==> (@Oeq _)
|
||||
as Forget_eq_compat.
|
||||
intros d0 d1 dle. split ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma Forget_leinj: forall (D:cpo) (P:D -> Prop) (I:admissible P) (d:SubCPO I) e, Forget I d <= Forget I e -> d <= e.
|
||||
intros D P I d e. case e. clear e. intros e Pe. case d. clear d. intros d Pd.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Hint Resolve Forget_leinj.
|
||||
|
||||
Lemma Forget_inj : forall (D:cpo) (P:D -> Prop) (I:admissible P) (d:SubCPO I) e, Forget I d == Forget I e -> d == e.
|
||||
intros. split ; auto.
|
||||
Qed.
|
||||
|
||||
Hint Resolve Forget_inj.
|
||||
|
||||
Lemma InheritFun_eq_compat D E Q Qcc f g X XX : Oeq f g ->
|
||||
(@InheritFun D E Q Qcc f X) == (@InheritFun D E Q Qcc g XX).
|
||||
intros. refine (fcont_eq_intro _). intros d. auto.
|
||||
Qed.
|
||||
|
||||
Lemma ForgetInherit (D E:cpo) (P:D -> Prop) PS f B : Forget PS << @InheritFun E D P PS f B == f.
|
||||
intros D E P PS f B.
|
||||
refine (fcont_eq_intro _). intros d.
|
||||
rewrite fcont_comp_simpl. rewrite InheritFun_simpl.
|
||||
simpl. auto.
|
||||
Qed.
|
||||
|
||||
Lemma InheritFun_comp: forall D E F P I (f:E -C-> F) X (g:D -C-> E) XX,
|
||||
@InheritFun _ F P I f X << g == @InheritFun _ F P I (f << g) XX.
|
||||
intros. refine (fcont_eq_intro _). intros d. repeat (rewrite fcont_comp_simpl).
|
||||
split ; auto.
|
||||
Qed.
|
983
papers/domains-2009/PredomKleisli.v
Executable file
983
papers/domains-2009/PredomKleisli.v
Executable file
|
@ -0,0 +1,983 @@
|
|||
Require Import utility.
|
||||
Require Import PredomCore.
|
||||
Require Import PredomProd.
|
||||
Require Import PredomLift.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
Definition kleislit (D0 D1 : cpo) (f: D0 -> D1 _BOT) :=
|
||||
cofix kl (l:DL_ord D0) := match l with Eps l => Eps (kl l) | Val d => f d end.
|
||||
|
||||
Lemma kleisli_inv : forall (D0 D1:cpo) (f: D0 -> D1 _BOT) (l : D0 _BOT), kleislit f l =
|
||||
match l with Eps l' => Eps (kleislit f l')
|
||||
| Val d => f d
|
||||
end.
|
||||
intros.
|
||||
transitivity (match kleislit f l with Eps x => Eps x | Val d => Val d end).
|
||||
apply (DL_inv (D:=D1)).
|
||||
case l.
|
||||
auto.
|
||||
intro d.
|
||||
simpl.
|
||||
symmetry.
|
||||
apply DL_inv.
|
||||
Qed.
|
||||
|
||||
(* Monad stuff *)
|
||||
Lemma left_unit : forall (D0 D1 : cpo) (d:D0) (f : D0 -c> DL D1), kleislit f (Val d) == f d.
|
||||
intros; auto. rewrite kleisli_inv.
|
||||
trivial.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_Eps: forall D E (x : DL D) (f:D -> DL E), kleislit f (Eps x) = Eps (kleislit f x).
|
||||
intros.
|
||||
rewrite kleisli_inv. auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_Val: forall (D:cpo) E (d:D) (f:D -> DL E), kleislit f (Val d) = f d.
|
||||
Proof.
|
||||
intros.
|
||||
rewrite kleisli_inv.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
(* Revising kleislival to work with nats first cos that'll make the induction work better... *)
|
||||
Lemma kleislipredValVal : forall (D E : cpo) k (e: tcpo E) (M: DL_ord D) (f : D -> DL E), pred_nth (kleislit f M) k = Val e -> exists d, M == Val d /\ f d == Val e.
|
||||
induction k.
|
||||
|
||||
intros.
|
||||
simpl in H.
|
||||
rewrite kleisli_inv in H.
|
||||
dcase M; intros.
|
||||
rewrite H0 in H.
|
||||
discriminate.
|
||||
rewrite H0 in H.
|
||||
exists t.
|
||||
split.
|
||||
trivial.
|
||||
rewrite H; auto.
|
||||
|
||||
intros.
|
||||
rewrite kleisli_inv in H.
|
||||
dcase M; intros.
|
||||
rewrite H0 in H.
|
||||
simpl pred_nth in H.
|
||||
assert (Hfoo := IHk _ _ _ H).
|
||||
destruct Hfoo as (d0,(H1,H2)).
|
||||
exists d0.
|
||||
split.
|
||||
rewrite <- H1.
|
||||
apply Oeq_sym.
|
||||
apply eq_Eps.
|
||||
assumption.
|
||||
|
||||
rewrite H0 in H.
|
||||
exists t.
|
||||
split.
|
||||
auto.
|
||||
rewrite <- H.
|
||||
(* this is bloody obvious, why has auto foresaken me? *)
|
||||
Lemma predn_DLeq : forall D n (x:DL_ord D), x == pred_nth x n.
|
||||
induction n; intros.
|
||||
simpl; auto.
|
||||
rewrite pred_nth_Sn.
|
||||
set (GRR := pred_nth x n).
|
||||
rewrite (IHn x).
|
||||
subst GRR.
|
||||
apply predeq.
|
||||
Qed.
|
||||
|
||||
apply (@predn_DLeq E (S k) (f t)).
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_Valeq: forall (D E:cpo) v (d:D) (f : D -m> E _BOT), v == Val d ->
|
||||
kleislit f v == (f d:DL_ord (E)).
|
||||
Proof.
|
||||
intros D E v d M veq.
|
||||
destruct veq as [v1 v2].
|
||||
destruct (DLle_Val_exists_pred v2) as [n [dd [pd ld]]].
|
||||
assert (dd == d) as DD. apply Ole_antisym.
|
||||
assert (xx:=DLle_pred_nth_left). specialize (xx _ n _ _ v1). rewrite pd in xx.
|
||||
apply vleinj. apply xx. apply ld. clear ld.
|
||||
|
||||
assert (M d == M dd). assert (stable M) by auto. unfold stable in H.
|
||||
apply H. auto. rewrite H. clear d DD H v1 v2.
|
||||
|
||||
generalize dependent v.
|
||||
induction n.
|
||||
intros v P. simpl in P. rewrite P in *.
|
||||
rewrite kleisli_Val. auto.
|
||||
|
||||
intros v P. destruct v.
|
||||
rewrite kleisli_Eps.
|
||||
simpl in P.
|
||||
apply Oeq_trans with (kleislit M v). apply (Oeq_sym (@eq_Eps E (kleislit M v))).
|
||||
rewrite (IHn _ P) ; auto.
|
||||
|
||||
inversion P as [Eq]. rewrite Eq in *. simpl in P.
|
||||
assert (pred_nth (Val dd) n = Val dd). destruct n ; auto.
|
||||
apply (IHn _ H).
|
||||
Qed.
|
||||
|
||||
Hint Resolve DLle_leVal.
|
||||
|
||||
Lemma kleisliValVal : forall (D E : cpo) (M: D _BOT) (f : D -c> E _BOT) e,
|
||||
kleislit f M == Val e -> exists d, M == Val d /\ f d == Val e.
|
||||
intros D E M N e [kv1 kv2].
|
||||
destruct (DLvalval kv2) as [n [d [pd ld]]].
|
||||
destruct (kleislipredValVal pd) as [md [vm ndm]].
|
||||
exists md. split. apply vm. refine (Oeq_trans ndm _).
|
||||
rewrite (kleisli_Valeq (fcontit N) vm) in kv1. rewrite ndm in kv1.
|
||||
split. apply kv1.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Definition kleislim (D E:cpo) : (D-c> DL E) -> (DL_ord D) -m> DL E.
|
||||
intros.
|
||||
exists (kleislit X).
|
||||
unfold monotonic.
|
||||
intros x y L. simpl. apply DLless_cond.
|
||||
intros e C. destruct (kleisliValVal C) as [dd [P Q]].
|
||||
rewrite (kleisli_Valeq (fcontit X) P).
|
||||
rewrite P in L.
|
||||
destruct (DLle_Val_exists_eq L) as [y0 [yv Py]].
|
||||
rewrite (kleisli_Valeq (fcontit X) yv).
|
||||
assert (monotonic (fcontit X)) as M by auto.
|
||||
apply M. auto.
|
||||
Defined.
|
||||
|
||||
Lemma eta_val: forall D f, @eta D f = Val f.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Definition kleisli (D0 D1: cpo) : (D0 -c> DL D1) -> (DL D0) -c> DL D1.
|
||||
intros D0 D1 f.
|
||||
exists (kleislim f).
|
||||
unfold continuous.
|
||||
intros h.
|
||||
simpl.
|
||||
apply DLless_cond.
|
||||
intros xx C.
|
||||
destruct (kleisliValVal C) as [x' [P Q]].
|
||||
rewrite (kleisli_Valeq (fcontit f) P).
|
||||
apply Ole_trans with (y:=lub (kleislim f <m< h)) ; auto.
|
||||
clear C Q xx.
|
||||
destruct (lubval P) as [k [hk [hkv Pk]]].
|
||||
destruct (DLvalgetchain hkv) as [c Pc].
|
||||
apply Ole_trans with (y:=((lub (fcontit f <m< c)))).
|
||||
assert (continuous (fcontit f)) as cf by auto.
|
||||
rewrite <- (lub_comp_eq c cf).
|
||||
assert (monotonic f) as M by auto. apply M.
|
||||
apply vleinj. rewrite <- P.
|
||||
rewrite <- eta_val.
|
||||
assert (continuous (fcontit (@eta D0))) as ce by auto.
|
||||
rewrite (lub_comp_eq c ce).
|
||||
apply Ole_trans with (y:= lub h) ; auto.
|
||||
rewrite (lub_lift_left _ k). refine (lub_le_compat _).
|
||||
apply fmon_le_intro. intros n. destruct (Pc n) ; auto.
|
||||
rewrite (lub_lift_left (kleislim f <m< h) k).
|
||||
refine (lub_le_compat _). apply fmon_le_intro.
|
||||
intros n. repeat (rewrite fmon_comp_simpl).
|
||||
apply Ole_trans with (y:=kleislit f (h (k+n))) ; auto.
|
||||
specialize (Pc n). rewrite (kleisli_Valeq (fcontit f) Pc).
|
||||
auto.
|
||||
Defined.
|
||||
|
||||
Lemma kleisli_simpl (D0 D1: cpo) (f:D0 -c> D1 _BOT) v : kleisli f v = kleislit f v.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Definition Kleisli (D0 D1 : cpo) : (D0 -C-> DL D1) -m> (DL D0) -C-> DL D1.
|
||||
intros. exists (@kleisli D0 D1).
|
||||
unfold monotonic.
|
||||
intros f g fgL x.
|
||||
simpl. apply DLless_cond.
|
||||
intros e C. destruct (kleisliValVal C) as [d [P Q]].
|
||||
rewrite (kleisli_Valeq (fcontit f) P).
|
||||
rewrite (kleisli_Valeq (fcontit g) P). apply fgL.
|
||||
Defined.
|
||||
|
||||
Lemma Kleisli_simpl (D0 D1 : cpo) (f:D0 -c> D1 _BOT) : Kleisli _ _ f = kleisli f.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Definition KLEISLI (D0 D1: cpo) : (D0 -C-> DL D1) -c> DL D0 -C-> DL D1.
|
||||
intros D0 D1. exists (Kleisli D0 D1).
|
||||
unfold continuous.
|
||||
intros c. intros d0. simpl.
|
||||
apply DLless_cond.
|
||||
intros e C.
|
||||
destruct (kleisliValVal C) as [d [V hd]].
|
||||
rewrite (@kleisli_Valeq _ _ _ _ (fcontit (fcont_lub c)) V).
|
||||
apply Ole_trans with (y:=lub (c <__> d)) ; auto.
|
||||
apply Ole_trans with (y:=lub (Kleisli D0 D1 <m< c <__> d0)) ; auto.
|
||||
refine (lub_le_compat _). apply fmon_le_intro. intros n.
|
||||
repeat (rewrite fcont_app_simpl). rewrite fmon_comp_simpl.
|
||||
rewrite Kleisli_simpl.
|
||||
rewrite kleisli_simpl.
|
||||
rewrite (kleisli_Valeq (fcontit (c n)) V). auto.
|
||||
Defined.
|
||||
|
||||
Add Parametric Morphism (D0 D1:cpo) f : (@kleisli D0 D1 f)
|
||||
with signature (@Ole (D0 _BOT)) ++> (@Ole (D1 _BOT))
|
||||
as kleisli_le_compat.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D0 D1:cpo) f : (@kleisli D0 D1 f)
|
||||
with signature (@Oeq (D0 _BOT)) ==> (@Oeq (D1 _BOT))
|
||||
as kleisli_eq_compat.
|
||||
intros ; split ; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D0 D1:cpo) : (KLEISLI D0 D1)
|
||||
with signature (@Ole (D0 -c> D1 _BOT)) ++> (@Ole (D0 _BOT -C-> D1 _BOT))
|
||||
as KLEISLI_le_compat.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D0 D1:cpo) : (KLEISLI D0 D1)
|
||||
with signature (@Oeq (D0 -c> D1 _BOT)) ==> (@Oeq (D0 _BOT -C-> D1 _BOT))
|
||||
as KLEISLI_eq_compat.
|
||||
intros ; split ; auto.
|
||||
Qed.
|
||||
|
||||
Implicit Arguments KLEISLI [D0 D1].
|
||||
|
||||
Lemma KLEISLI_simpl (D0 D1: cpo) (f:D0 -c> D1 _BOT) : KLEISLI f = kleisli f.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleislit_comp: forall D E F (f:DL E -c> DL F) (g:D -c> DL E) d,
|
||||
(forall x xx, f x == Val xx -> exists y, x == Val y) ->
|
||||
kleislit (f << g) d == f (kleislit g d).
|
||||
intros D E F f g d S. split. simpl.
|
||||
apply DLless_cond.
|
||||
|
||||
intros ff kv.
|
||||
destruct (kleisliValVal kv) as [dd [dv fdd]].
|
||||
rewrite (kleisli_Valeq (fcontit (f << g)) dv). rewrite fdd.
|
||||
assert (f (g dd) == f (kleisli g d)).
|
||||
assert (stable f) as sf. auto.
|
||||
unfold stable in sf. apply sf.
|
||||
rewrite (kleisli_Valeq (fcontit g) dv). auto.
|
||||
rewrite <- H. auto.
|
||||
|
||||
simpl. apply DLless_cond.
|
||||
intros xx fx.
|
||||
specialize (S _ _ fx).
|
||||
destruct S as [e ke].
|
||||
destruct (kleisliValVal ke) as [dd [de ge]].
|
||||
rewrite (kleisli_Valeq (fcontit (f << g)) de).
|
||||
assert (monotonic f) as M by auto.
|
||||
apply Ole_trans with (y:= (f << g) dd) ; auto.
|
||||
rewrite fcont_comp_simpl.
|
||||
apply M. rewrite <- ge in ke.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_comp: forall D0 D1 D2 (f:D0 -c> DL D1) (g:D1 -c> DL D2),
|
||||
kleisli g << kleisli f == kleisli (kleisli g << f).
|
||||
intros.
|
||||
refine (fcont_eq_intro _).
|
||||
intros d. rewrite fcont_comp_simpl.
|
||||
repeat (rewrite KLEISLI_simpl).
|
||||
repeat (rewrite kleisli_simpl).
|
||||
rewrite kleislit_comp. auto.
|
||||
intros e ff. rewrite kleisli_simpl.
|
||||
intros kl. destruct (kleisliValVal kl) as [dd [edd gd]].
|
||||
exists dd. auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_leq: forall (D E F:cpo) a b (f:D -C-> DL E) (g: F -C-> DL E),
|
||||
(forall aa, a == Val aa -> exists bb, b == Val bb /\ f aa <= g bb) ->
|
||||
@kleisli D E f a <= @kleisli F E g b.
|
||||
Proof.
|
||||
intros D E F a b f g V. simpl. apply DLless_cond.
|
||||
intros xx [_ kl].
|
||||
destruct (DLle_Val_exists_pred kl) as [n [dd [pd ld]]].
|
||||
destruct (kleislipredValVal pd) as [aa [va fa]].
|
||||
rewrite (kleisli_Valeq (fcontit f) va). specialize (V _ va).
|
||||
destruct V as [bb [vb fg]]. rewrite (kleisli_Valeq (fcontit g) vb). apply fg.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_eq: forall (D E F:cpo) a b (f:D -C-> DL E) (g: F -C-> DL E),
|
||||
(forall aa, a == Val aa -> exists bb, b == Val bb /\ f aa == g bb) ->
|
||||
(forall bb, b == Val bb -> exists aa, a == Val aa /\ f aa == g bb) ->
|
||||
@kleisli D E f a == @kleisli F E g b.
|
||||
Proof.
|
||||
intros D E F a b f g V1 V2.
|
||||
apply Ole_antisym.
|
||||
apply kleisli_leq.
|
||||
intros aa va. destruct (V1 _ va) as [bb [vb fa]].
|
||||
exists bb. auto.
|
||||
apply kleisli_leq. intros bb vb.
|
||||
destruct (V2 _ vb) as [aa [f1 f2]].
|
||||
exists aa. auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_eta_com: forall D E f, @kleisli D E f << eta == f.
|
||||
intros D E f. refine (fcont_eq_intro _).
|
||||
intros d. rewrite fcont_comp_simpl.
|
||||
rewrite eta_val. rewrite kleisli_simpl. rewrite kleisli_Val. auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_point_unit: forall (D:cpo) (d:D _BOT), kleisli eta d == d.
|
||||
intros D d.
|
||||
split. simpl. apply DLless_cond.
|
||||
|
||||
intros dd kl. destruct (kleisliValVal kl) as [d1 [P1 P2]].
|
||||
assert (eta d1 == Val dd) as P3 by auto. clear P2.
|
||||
rewrite eta_val in P3. rewrite P3 in P1.
|
||||
rewrite (kleisli_Valeq (fcontit eta) P1). auto.
|
||||
|
||||
simpl. apply DLless_cond.
|
||||
intros dd kl.
|
||||
rewrite (kleisli_Valeq (fcontit eta) kl). auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_unit: forall D, kleisli (@eta D) == ID.
|
||||
intros D. refine (fcont_eq_intro _).
|
||||
intros d. rewrite ID_simpl. simpl.
|
||||
assert (xx := kleisli_point_unit). simpl. simpl in xx. apply xx.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_comp2: forall D E F (f:D -C-> E) (g:E -C-> DL F),
|
||||
kleisli (g << f) == kleisli g << kleisli (eta << f).
|
||||
intros D E F f g.
|
||||
refine (fcont_eq_intro _).
|
||||
intros d. repeat (rewrite kleisli_cc_simpl).
|
||||
rewrite fcont_comp_simpl. repeat (rewrite kleisli_cc_simpl).
|
||||
refine (kleisli_eq _ _).
|
||||
intros dd dv. exists (f dd). rewrite fcont_comp_simpl. split ; auto.
|
||||
rewrite (kleisli_Valeq (fcontit (eta << f)) dv). auto.
|
||||
intros ee kl.
|
||||
destruct (kleisliValVal kl) as [dd P]. exists dd.
|
||||
split. destruct P ; auto.
|
||||
rewrite fcont_comp_simpl in *. destruct P as [_ P].
|
||||
simpl in P.
|
||||
assert (stable g) as sg by auto. apply sg.
|
||||
apply (vinj P).
|
||||
Qed.
|
||||
|
||||
Definition mu D := kleisli (@ID (D _BOT)).
|
||||
|
||||
Lemma mu_natural D E (f:D -c> E) :
|
||||
kleisli (eta << f) << mu _ == mu _ << kleisli (eta << (kleisli (eta << f))).
|
||||
intros D E f.
|
||||
unfold mu. rewrite <- kleisli_comp2. rewrite kleisli_comp.
|
||||
rewrite fcont_comp_unitR. rewrite fcont_comp_unitL.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma mumu D : mu _ << kleisli (eta << mu _) == mu D << mu _.
|
||||
intros D. unfold mu. rewrite <- kleisli_comp2. rewrite kleisli_comp.
|
||||
repeat (rewrite fcont_comp_unitL). rewrite fcont_comp_unitR. auto.
|
||||
Qed.
|
||||
|
||||
Lemma mukl D : mu D << kleisli (eta << eta) == ID.
|
||||
intros D. unfold mu.
|
||||
rewrite <- kleisli_comp2. rewrite fcont_comp_unitL. rewrite kleisli_unit. auto.
|
||||
Qed.
|
||||
|
||||
|
||||
(* Now that I've got bottomless cpos, I need ones with bottom for fixpoints.
|
||||
Could work with things that are lifted, but that's a bit literal.
|
||||
Could work with algebras for the lift monad, but that's probably going to be heavy in Coq
|
||||
So I'll just add a new class of things with bottom, like Christine's original cpos but with
|
||||
a new layer of stuff. Hope I can do enough implicit stuff not to make this too disgusting.
|
||||
*)
|
||||
(*
|
||||
Record pcpo : Type := mk_pcpo
|
||||
{tpcpo:>cpo; D0 : tpcpo; Dbot : forall x:tpcpo, D0 <= x}.
|
||||
|
||||
Check DL.
|
||||
|
||||
Check DL_bot.
|
||||
|
||||
Definition Lift (D:cpo) : pcpo.
|
||||
intro D; exists (DL D) (DL_bot D).
|
||||
intro.
|
||||
red. simpl.
|
||||
apply (DL_bot_least).
|
||||
Defined.
|
||||
*)
|
||||
|
||||
Record alg : Type := mk_alg
|
||||
{ acpo :> cpo; strength : DL acpo -c> acpo;
|
||||
st_eta : strength << eta == ID;
|
||||
st_kl : strength << kleisli (eta << strength) == strength << kleisli ID }.
|
||||
|
||||
Definition Aprod : alg -> alg -> alg.
|
||||
intros A1 A2. exists (Dprod (acpo A1) (acpo A2))
|
||||
(<| strength A1 << (kleisli (eta << FST)),
|
||||
strength A2 << (kleisli (eta << SND)) |>).
|
||||
rewrite PROD_fun_compl.
|
||||
apply (@Dprod_unique A1 A2 (Dprod A1 A2)). rewrite FST_PROD_fun.
|
||||
rewrite fcont_comp_unitR. rewrite fcont_comp_assoc.
|
||||
rewrite kleisli_eta_com. rewrite <- fcont_comp_assoc. rewrite st_eta. rewrite fcont_comp_unitL. auto.
|
||||
rewrite SND_PROD_fun.
|
||||
rewrite fcont_comp_unitR. rewrite fcont_comp_assoc.
|
||||
rewrite kleisli_eta_com. rewrite <- fcont_comp_assoc. rewrite st_eta. rewrite fcont_comp_unitL. auto.
|
||||
|
||||
refine (@Dprod_unique (acpo A1) (acpo A2) _ _ _ _ _). rewrite <- fcont_comp_assoc.
|
||||
rewrite FST_PROD_fun. rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun.
|
||||
rewrite fcont_comp_assoc. rewrite fcont_comp_assoc.
|
||||
rewrite kleisli_comp. rewrite <- fcont_comp_assoc. rewrite kleisli_eta_com.
|
||||
rewrite fcont_comp_assoc. rewrite FST_PROD_fun. rewrite kleisli_comp.
|
||||
rewrite fcont_comp_unitR.
|
||||
rewrite <- fcont_comp_assoc. rewrite kleisli_comp2. rewrite <- fcont_comp_assoc.
|
||||
rewrite st_kl. rewrite fcont_comp_assoc. rewrite <- kleisli_comp2. rewrite fcont_comp_unitL.
|
||||
auto.
|
||||
|
||||
rewrite <- fcont_comp_assoc.
|
||||
rewrite SND_PROD_fun. rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun.
|
||||
rewrite fcont_comp_assoc. rewrite fcont_comp_assoc.
|
||||
rewrite kleisli_comp. rewrite <- fcont_comp_assoc. rewrite kleisli_eta_com.
|
||||
rewrite fcont_comp_assoc. rewrite SND_PROD_fun. rewrite kleisli_comp.
|
||||
rewrite fcont_comp_unitR.
|
||||
rewrite <- fcont_comp_assoc. rewrite kleisli_comp2. rewrite <- fcont_comp_assoc.
|
||||
rewrite st_kl. rewrite fcont_comp_assoc. rewrite <- kleisli_comp2. rewrite fcont_comp_unitL.
|
||||
auto.
|
||||
Defined.
|
||||
|
||||
Definition AF : cpo -> alg.
|
||||
intros D. exists (DL D) (kleisli ID).
|
||||
rewrite kleisli_eta_com. auto.
|
||||
rewrite <- kleisli_comp2. rewrite fcont_comp_unitL.
|
||||
rewrite kleisli_comp. rewrite fcont_comp_unitR. auto.
|
||||
Defined.
|
||||
|
||||
Definition Application := fun D0 D1 => curry (uncurry (@KLEISLI (D0 -C-> D1 _BOT) D1) <<
|
||||
<| curry (uncurry KLEISLI << SWAP) << SND, FST |>).
|
||||
|
||||
Definition Operator2 := fun A B C (F:Dprod A B -C-> DL C) =>
|
||||
(Application _ _) << (kleisli (eta << (curry F))).
|
||||
|
||||
Definition Smash := fun A B => @Operator2 _ _ (Dprod A B) eta.
|
||||
|
||||
Definition LStrength := fun A B => uncurry (Smash A _) << eta *f* (@ID (DL B)).
|
||||
Definition RStrength := fun A B => uncurry (Smash _ B) << (@ID (DL A)) *f* eta.
|
||||
|
||||
Definition KLEISLIR := fun A B C (f:Dprod A B -c> DL C) => kleisli f << LStrength A B.
|
||||
Definition KLEISLIL := fun A B C (f:Dprod A B -c> DL C) => kleisli f << RStrength A B.
|
||||
|
||||
Lemma Operator2Val: forall (D E F:cpo) (h:Dprod D E -C-> DL F) d e f,
|
||||
Operator2 h d e == Val f -> exists d', exists e', d == Val d' /\ e == Val e' /\ h (d',e') == Val f.
|
||||
intros D E F h d e f.
|
||||
unfold Operator2. unfold Application.
|
||||
repeat (try (rewrite fcont_comp_simpl) ; try (rewrite curry_simpl) ; try (rewrite kliesli_cc_simpl) ;
|
||||
try (rewrite PROD_fun_simpl) ; simpl ; try (rewrite FST_simpl) ; try (rewrite SND_simpl) ;
|
||||
try (rewrite fcont_comp2_simpl) ; try (rewrite kleisli_c_simpl) ;
|
||||
try (rewrite ID_simpl)).
|
||||
rewrite uncurry_simpl. rewrite KLEISLI_simpl. rewrite kleisli_simpl.
|
||||
intros c.
|
||||
destruct (kleisliValVal c) as [f' [cc1 cc2]]. clear c. rewrite curry_simpl in cc2.
|
||||
rewrite fcont_comp_simpl in cc2. unfold SWAP in cc2. rewrite PROD_fun_simpl in cc2.
|
||||
rewrite uncurry_simpl in cc2. rewrite SND_simpl, FST_simpl in cc2. simpl in cc2. rewrite KLEISLI_simpl in cc2.
|
||||
destruct (kleisliValVal cc2) as [e' [eeq feeq]]. clear cc2.
|
||||
destruct (kleisliValVal cc1) as [d' [deq deeq]]. clear cc1.
|
||||
rewrite fcont_comp_simpl in *.
|
||||
exists d'. exists e'. split. auto. split. auto.
|
||||
rewrite eta_val in *. assert (xx:=fcont_eq_elim (vinj deeq) e'). clear deeq deq eeq. rewrite <- feeq.
|
||||
rewrite curry_simpl in xx. auto.
|
||||
Qed.
|
||||
|
||||
Lemma Operator2_simpl: forall (E F D:cpo) (f:Dprod E F -C-> DL D) v1 v2,
|
||||
Operator2 f (Val v1) (Val v2) == f (v1,v2).
|
||||
intros E F D f.
|
||||
intros v1 v2.
|
||||
unfold Operator2. unfold Application. repeat (rewrite fcont_comp_simpl).
|
||||
rewrite kleisli_simpl. rewrite kleisli_Val. rewrite curry_simpl.
|
||||
rewrite fcont_comp_simpl. rewrite PROD_fun_simpl. rewrite uncurry_simpl.
|
||||
rewrite FST_simpl. simpl. rewrite KLEISLI_simpl.
|
||||
rewrite kleisli_simpl. repeat (rewrite fcont_comp_simpl). rewrite eta_val. rewrite kleisli_Val.
|
||||
rewrite SND_simpl. simpl. rewrite curry_simpl. rewrite fcont_comp_simpl.
|
||||
unfold SWAP. rewrite PROD_fun_simpl. rewrite SND_simpl, FST_simpl. simpl.
|
||||
rewrite uncurry_simpl. rewrite KLEISLI_simpl. rewrite kleisli_simpl. rewrite kleisli_Val.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma strength_proj D E: kleisli (eta << SND) << LStrength _ _ == @SND E (DL D).
|
||||
intros D E.
|
||||
unfold LStrength. unfold Smash. refine (fcont_eq_intro _).
|
||||
intros p. case p. clear p. simpl. intros u d.
|
||||
rewrite SND_simpl. simpl. repeat (rewrite fcont_comp_simpl).
|
||||
rewrite PROD_map_simpl. simpl. rewrite ID_simpl. simpl.
|
||||
rewrite uncurry_simpl. split.
|
||||
refine (DLless_cond _). intros dd C.
|
||||
destruct (kleisliValVal C) as [x [P Pd]].
|
||||
rewrite fcont_comp_simpl in *. rewrite eta_val in Pd. assert (Pdd:=vinj Pd). clear Pd.
|
||||
destruct x as [uu x]. rewrite SND_simpl in *. simpl in Pdd.
|
||||
destruct (Operator2Val P) as [u1 [d1 [_ Pu]]]. clear P C. rewrite eta_val in Pu.
|
||||
assert (P := vinj (proj2 Pu)). repeat (rewrite eta_val).
|
||||
assert (Pd := proj1 Pu).
|
||||
apply Ole_trans with (y:=kleisli (eta << SND) (Operator2 (eta) (Val (D:=E) u) (Val d1))) ; auto.
|
||||
refine (app_le_compat (Ole_refl _) _). refine (app_le_compat (Ole_refl _) _) ; auto.
|
||||
rewrite (app_eq_compat (Oeq_refl (kleisli (eta << SND))) (Operator2_simpl (@eta (Dprod E D)) _ _)).
|
||||
repeat (rewrite eta_val). rewrite kleisli_simpl, kleisli_Val. auto.
|
||||
|
||||
refine (DLless_cond _). intros dd V.
|
||||
rewrite <- (app_eq_compat (Oeq_refl (KLEISLI (eta << SND)))
|
||||
(app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod E D)) (@eta E u))) (Oeq_sym V))).
|
||||
rewrite eta_val. rewrite (app_eq_compat (Oeq_refl (KLEISLI (@eta D << @SND E D)))
|
||||
(Operator2_simpl (@eta (Dprod E D)) _ _)).
|
||||
rewrite KLEISLI_simpl. rewrite eta_val. rewrite kleisli_simpl, kleisli_Val. auto.
|
||||
Qed.
|
||||
|
||||
Definition assoc D E F := (PROD_fun (FST << FST) (PROD_fun (@SND D E << FST) (@SND _ F))).
|
||||
Lemma assoc_simpl D E F d e f : assoc D E F (d,e,f) = (d,(e,f)).
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma strength_assoc D E F :
|
||||
LStrength _ _ << <| FST , LStrength _ _ << SND |> << assoc D E (DL F) ==
|
||||
kleisli (eta << assoc _ _ _) << LStrength _ _.
|
||||
intros D E F. refine (fcont_eq_intro _).
|
||||
intros p. case p ; clear p. intros p df. case p ; clear p ; intros d e.
|
||||
repeat (rewrite fcont_comp_simpl). unfold LStrength. repeat (rewrite fcont_comp_simpl).
|
||||
rewrite assoc_simpl. rewrite PROD_fun_simpl. repeat (rewrite fcont_comp_simpl).
|
||||
rewrite FST_simpl. simpl. unfold PROD_map.
|
||||
repeat (rewrite PROD_fun_simpl). repeat (rewrite fcont_comp_simpl). rewrite FST_simpl, SND_simpl, ID_simpl. simpl.
|
||||
rewrite ID_simpl, SND_simpl. simpl. repeat (rewrite uncurry_simpl).
|
||||
repeat (rewrite FST_simpl, SND_simpl). rewrite ID_simpl. simpl.
|
||||
unfold Smash. split.
|
||||
refine (DLless_cond _). intros p C.
|
||||
destruct (Operator2Val C) as [dd [ee [ed [Ov Pd]]]].
|
||||
rewrite eta_val in Pd. assert (PPd := vinj Pd). clear Pd. rewrite eta_val in ed.
|
||||
assert (Ed := vinj ed). clear ed. rewrite <- Ed in PPd. clear dd Ed.
|
||||
destruct ee as [ee ff].
|
||||
destruct (Operator2Val Ov) as [e1 [f1 [Eeq [deq Pe]]]].
|
||||
rewrite eta_val in *. clear Ov. assert (EEq := vinj Eeq). clear Eeq.
|
||||
assert (PPe := vinj Pe). clear Pe C.
|
||||
rewrite (app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod D (Dprod E F))) (Val (D:=D) d)))
|
||||
(app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod E F)) (@eta E e))) deq)).
|
||||
rewrite eta_val.
|
||||
rewrite (app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod D (Dprod E F))) (Val (D:=D) d)))(Operator2_simpl _ _ _)).
|
||||
repeat (rewrite eta_val). rewrite Operator2_simpl.
|
||||
rewrite (app_eq_compat (Oeq_refl (kleisli (@eta (Dprod D (Dprod E F)) << assoc D E F))) (app_eq_compat (Oeq_refl _) deq)).
|
||||
rewrite (app_eq_compat (Oeq_refl (kleisli (@eta (Dprod D (Dprod E F)) << assoc D E F)))
|
||||
(Operator2_simpl (@eta (Dprod (Dprod D E) F)) _ _)). rewrite kleisli_simpl.
|
||||
repeat (rewrite eta_val). rewrite kleisli_Val. rewrite fcont_comp_simpl. rewrite assoc_simpl.
|
||||
auto.
|
||||
|
||||
refine (DLless_cond _).
|
||||
intros p C. rewrite kleisli_simpl in C. destruct (kleisliValVal C) as [pp [OV CC]]. clear C.
|
||||
destruct (Operator2Val OV) as [p1 [f [fv [def Px]]]].
|
||||
rewrite eta_val in Px. assert (ppeq := vinj Px). clear Px.
|
||||
rewrite eta_val in fv. assert (deq := vinj fv). clear fv. clear OV.
|
||||
rewrite (app_eq_compat (Oeq_refl (kleisli (@eta (Dprod D (Dprod E F)) << assoc D E F)))
|
||||
(app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod (Dprod D E) F)) (@eta (Dprod D E) (d, e)))) def)).
|
||||
rewrite (app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod D (Dprod E F))) (@eta D d)))
|
||||
(app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod E F)) (@eta E e))) def)).
|
||||
repeat (rewrite eta_val).
|
||||
rewrite (app_eq_compat (Oeq_refl (kleisli (@eta (Dprod D (Dprod E F)) << assoc D E F))) (Operator2_simpl (@eta (Dprod (Dprod D E) F)) _ _)).
|
||||
repeat (rewrite eta_val).
|
||||
rewrite kleisli_simpl, kleisli_Val. rewrite fcont_comp_simpl.
|
||||
rewrite (app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod D (Dprod E F))) (Val (D:=D) d))) (Operator2_simpl (@eta (Dprod E F)) _ _)).
|
||||
repeat (rewrite eta_val). rewrite Operator2_simpl.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma strength_eta D E : LStrength _ _ << PROD_fun (@FST D E) (eta << SND) == eta.
|
||||
intros D E.
|
||||
refine (fcont_eq_intro _).
|
||||
intros p. case p ; clear p.
|
||||
intros d e. repeat (rewrite fcont_comp_simpl). rewrite PROD_fun_simpl.
|
||||
rewrite fcont_comp_simpl.
|
||||
rewrite SND_simpl, FST_simpl. simpl.
|
||||
unfold LStrength. repeat (rewrite fcont_comp_simpl). unfold PROD_map. rewrite PROD_fun_simpl.
|
||||
repeat (rewrite fcont_comp_simpl). rewrite SND_simpl, FST_simpl. simpl. rewrite ID_simpl. simpl.
|
||||
rewrite uncurry_simpl. unfold Smash. repeat (rewrite eta_val). rewrite Operator2_simpl. auto.
|
||||
Qed.
|
||||
|
||||
Lemma strength_tt D E : LStrength _ _ << PROD_fun (@FST D (DL (DL E))) (mu _ << SND) ==
|
||||
mu _ << kleisli (eta << LStrength _ _) << LStrength _ _.
|
||||
intros D E. unfold mu.
|
||||
rewrite <- kleisli_comp2. rewrite fcont_comp_unitL.
|
||||
unfold LStrength. rewrite fcont_comp_assoc. rewrite <- PROD_fun_compr. unfold Smash.
|
||||
refine (fcont_eq_intro _).
|
||||
intros p ; case p ; clear p ; intros d e.
|
||||
repeat (rewrite fcont_comp_simpl).
|
||||
rewrite PROD_fun_simpl. repeat (rewrite fcont_comp_simpl). rewrite uncurry_simpl.
|
||||
repeat (rewrite ID_simpl). simpl. rewrite FST_simpl, SND_simpl. simpl.
|
||||
rewrite PROD_map_simpl. simpl. rewrite ID_simpl. rewrite uncurry_simpl. simpl.
|
||||
repeat (rewrite eta_val).
|
||||
split. refine (DLless_cond _). intros p. case p ; clear p ; intros dd ee.
|
||||
intros C. destruct (Operator2Val C) as [d1 [e1 [V1 [k Pe]]]].
|
||||
assert (dee := vinj V1). clear V1. assert (ddeq := vinj Pe). clear Pe.
|
||||
rewrite <- dee in ddeq. clear d1 dee. clear C. rewrite kleisli_simpl in k.
|
||||
destruct (kleisliValVal k) as [e2 [Pe Pee]]. rewrite ID_simpl in Pee. simpl in Pee.
|
||||
rewrite (app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod D E)) (Val d)))
|
||||
(app_eq_compat (Oeq_refl (kleisli ID)) Pe)).
|
||||
rewrite kleisli_simpl. rewrite kleisli_Val. rewrite ID_simpl.
|
||||
rewrite (app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod D E)) (Val d))) Pee). rewrite Operator2_simpl.
|
||||
rewrite (app_eq_compat (Oeq_refl (kleisli (uncurry (Operator2 (@eta (Dprod D E))) << PROD_map (eta) (ID))))
|
||||
(app_eq_compat (Oeq_refl (Operator2 (@eta (Dprod D (DL E))) (Val d))) Pe)).
|
||||
rewrite Operator2_simpl.
|
||||
repeat (rewrite eta_val). rewrite kleisli_simpl. rewrite kleisli_Val.
|
||||
rewrite fcont_comp_simpl, PROD_map_simpl. rewrite uncurry_simpl. rewrite ID_simpl.
|
||||
rewrite (app_eq_compat (Oeq_refl (Operator2 eta (eta (d)))) Pee).
|
||||
rewrite eta_val, Operator2_simpl. auto.
|
||||
|
||||
refine (DLless_cond _). intros p ; case p ; clear p ; intros dd ee.
|
||||
intros C. rewrite kleisli_simpl in C. destruct (kleisliValVal C) as [de [Ov Px]].
|
||||
destruct de as [d1 e1]. rewrite fcont_comp_simpl in Px. rewrite PROD_map_simpl, uncurry_simpl in Px.
|
||||
simpl in Ov,Px. rewrite ID_simpl in *. simpl in Px.
|
||||
destruct (Operator2Val Ov) as [d2 [e2 [deq [eeq Pv]]]].
|
||||
assert (ddeq := vinj deq). clear deq. clear Ov C.
|
||||
rewrite eta_val in Pv. assert (peq := vinj Pv). clear Pv.
|
||||
rewrite <- ddeq in peq. clear d2 ddeq.
|
||||
rewrite (app_eq_compat (Oeq_refl (Operator2 eta (Val d)))
|
||||
(app_eq_compat (Oeq_refl (kleisli ID)) eeq)).
|
||||
rewrite kleisli_simpl. rewrite kleisli_simpl, kleisli_Val. rewrite ID_simpl.
|
||||
destruct (Operator2Val Px) as [d3 [e3 [Pe [Pd Q]]]].
|
||||
assert (stable (@SND D (DL E))) as sf by auto. specialize (sf _ _ peq). repeat (rewrite SND_simpl in sf). simpl in sf.
|
||||
rewrite <- sf in *. clear sf e1 Px peq.
|
||||
rewrite (app_eq_compat (Oeq_refl (Operator2 eta (Val d))) Pd). rewrite Operator2_simpl.
|
||||
rewrite <- kleisli_simpl.
|
||||
rewrite (app_eq_compat (Oeq_refl (kleisli (uncurry (Operator2 eta) << eta *f* ID))) (app_eq_compat (Oeq_refl (Operator2 eta (Val d))) eeq)).
|
||||
rewrite Operator2_simpl. rewrite eta_val. rewrite kleisli_simpl.
|
||||
rewrite kleisli_Val. rewrite fcont_comp_simpl.
|
||||
rewrite PROD_map_simpl. rewrite uncurry_simpl. rewrite ID_simpl. rewrite eta_val.
|
||||
rewrite (app_eq_compat (Oeq_refl (Operator2 eta (Val d))) Pd).
|
||||
rewrite Operator2_simpl. auto.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_eq (A A' B B' C:cpo) : forall (f:Dprod A B -C-> DL C) (f':Dprod A' B' -C-> DL C)
|
||||
a b a' b',
|
||||
(forall aa aa', a == Val aa -> a' == Val aa' -> f (aa,b) == f' (aa',b')) ->
|
||||
(forall aa, a == Val aa -> exists aa', a' == Val aa') ->
|
||||
(forall aa', a' == Val aa' -> exists aa, a == Val aa) ->
|
||||
@KLEISLIL A B C f (a,b) == @KLEISLIL A' B' C f' (a',b').
|
||||
Proof.
|
||||
intros A A' B B' C f f' a b a' b' V1 V2 V3.
|
||||
unfold KLEISLIL.
|
||||
unfold RStrength. repeat (rewrite fcont_comp_simpl).
|
||||
repeat (rewrite PROD_map_simpl). simpl. repeat (rewrite ID_simpl). simpl.
|
||||
repeat (rewrite uncurry_simpl). unfold Smash.
|
||||
refine (kleisli_eq _ _).
|
||||
intros aa sa.
|
||||
destruct (Operator2Val sa) as [a1 [b1 [Pa1 [Pb1 pv]]]].
|
||||
destruct (V2 _ Pa1) as [aa' Paa].
|
||||
exists (aa', b'). specialize (V1 a1 aa' Pa1 Paa). rewrite <- V1.
|
||||
rewrite eta_val in *.
|
||||
split.
|
||||
refine (Oeq_trans (app_eq_compat (app_eq_compat (Oeq_refl (Operator2 eta)) Paa) (Oeq_refl _)) _).
|
||||
rewrite Operator2_simpl. auto.
|
||||
auto.
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
refine (Oeq_trans (Oeq_sym (vinj pv)) _).
|
||||
refine (pair_eq_compat (Oeq_refl _) _).
|
||||
apply (Oeq_sym (vinj Pb1)).
|
||||
|
||||
intros bb ob. destruct (Operator2Val ob) as [aa' [bb' [av P]]].
|
||||
specialize (V3 _ av). destruct V3 as [aa va].
|
||||
specialize (V1 _ _ va av).
|
||||
exists (aa,b).
|
||||
split. rewrite eta_val.
|
||||
refine (Oeq_trans (app_eq_compat (app_eq_compat (Oeq_refl (Operator2 eta)) va) (Oeq_refl (Val b))) _).
|
||||
rewrite Operator2_simpl. rewrite eta_val. auto.
|
||||
refine (Oeq_trans V1 _).
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
destruct P as [P1 P2]. rewrite eta_val in P2. refine (Oeq_trans _ (vinj P2)).
|
||||
rewrite eta_val in P1.
|
||||
refine (pair_eq_compat (Oeq_refl _) (vinj P1)).
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIR_eq (A A' B B' C:cpo) : forall (f:Dprod A B -C-> DL C) (f':Dprod A' B' -C-> DL C)
|
||||
a b a' b',
|
||||
(forall aa aa', b == Val aa -> b' == Val aa' -> f (a,aa) == f' (a',aa')) ->
|
||||
(forall aa, b == Val aa -> exists aa', b' == Val aa') ->
|
||||
(forall aa', b' == Val aa' -> exists aa, b == Val aa) ->
|
||||
@KLEISLIR A B C f (a,b) == @KLEISLIR A' B' C f' (a',b').
|
||||
Proof.
|
||||
intros A A' B B' C f f' a b a' b' V1 V2 V3.
|
||||
unfold KLEISLIR.
|
||||
unfold LStrength. repeat (rewrite fcont_comp_simpl).
|
||||
repeat (rewrite PROD_map_simpl). repeat (rewrite ID_simpl).
|
||||
repeat (rewrite uncurry_simpl). unfold Smash.
|
||||
refine (kleisli_eq _ _).
|
||||
intros aa sa.
|
||||
destruct (Operator2Val sa) as [a1 [b1 [Pa1 [Pb1 pv]]]].
|
||||
destruct (V2 _ Pb1) as [aa' Paa].
|
||||
exists (a',aa'). specialize (V1 _ aa' Pb1 Paa). rewrite <- V1.
|
||||
rewrite eta_val in *.
|
||||
split.
|
||||
refine (Oeq_trans (app_eq_compat (Oeq_refl (Operator2 eta (Val a'))) (Paa )) _).
|
||||
rewrite Operator2_simpl. auto.
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
refine (Oeq_trans (Oeq_sym (vinj pv)) _).
|
||||
refine (pair_eq_compat _ (Oeq_refl _)).
|
||||
apply (Oeq_sym (vinj Pa1)).
|
||||
|
||||
intros bb ob. destruct (Operator2Val ob) as [aa' [bb' [av [bv P]]]].
|
||||
specialize (V3 _ bv). destruct V3 as [aa va].
|
||||
specialize (V1 _ _ va bv).
|
||||
exists (a,aa).
|
||||
split. rewrite eta_val.
|
||||
refine (Oeq_trans (app_eq_compat (Oeq_refl (Operator2 eta (Val a))) va) _).
|
||||
rewrite Operator2_simpl. rewrite eta_val. auto.
|
||||
refine (Oeq_trans V1 _).
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
rewrite eta_val in P. refine (Oeq_trans _ (vinj P)).
|
||||
rewrite eta_val in av.
|
||||
refine (pair_eq_compat (vinj av) (Oeq_refl _)).
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E F:cpo) : (@KLEISLIL D E F)
|
||||
with signature (@Oeq (Dprod D E -c> DL F)) ==> (@Oeq (Dprod (DL D) E -c> DL F))
|
||||
as KLEISLIL_eq_compat.
|
||||
intros f0 f1 feq.
|
||||
apply fcont_eq_intro. intros de. simpl.
|
||||
unfold KLEISLIL. repeat (rewrite fcont_comp_simpl).
|
||||
repeat (rewrite kleisli_cc_simpl).
|
||||
refine (app_eq_compat _ (Oeq_refl _)). rewrite feq. auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E F:cpo) : (@KLEISLIR D E F)
|
||||
with signature (@Oeq (Dprod D E -c> DL F)) ==> (@Oeq (Dprod D (DL E) -c> DL F))
|
||||
as KLEISLIR_eq_compat.
|
||||
intros f0 f1 feq.
|
||||
apply fcont_eq_intro. intros de. simpl.
|
||||
unfold KLEISLIR. repeat (rewrite fcont_comp_simpl).
|
||||
repeat (rewrite kleisli_cc_simpl). refine (app_eq_compat _ (Oeq_refl _)).
|
||||
rewrite feq ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_comp: forall D E F G (f:Dprod D E -C-> DL F) g h,
|
||||
KLEISLIL f << PROD_fun g h == KLEISLIL (f << PROD_map ID h) << PROD_fun g (@ID G).
|
||||
intros.
|
||||
apply fcont_eq_intro.
|
||||
intros gg. repeat (rewrite fcont_comp_simpl).
|
||||
repeat (rewrite PROD_fun_simpl). rewrite ID_simpl.
|
||||
simpl. refine (KLEISLIL_eq _ _ _).
|
||||
intros d0 d1 vd0 vd1. rewrite fcont_comp_simpl. rewrite PROD_map_simpl. simpl. repeat (rewrite ID_simpl). simpl.
|
||||
rewrite vd0 in vd1. assert (vv:=vinj vd1). assert (stable f) as sf by auto. unfold stable in sf.
|
||||
apply (sf (d0,h gg) (d1,h gg)). auto.
|
||||
intros d0 vd0. exists d0. auto. intros aa vv. eexists ; apply vv.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIR_comp: forall D E F G (f:Dprod D E -C-> DL F) g h,
|
||||
KLEISLIR f << PROD_fun h g == KLEISLIR (f << PROD_map h ID) << PROD_fun (@ID G) g.
|
||||
intros.
|
||||
apply fcont_eq_intro.
|
||||
intros gg. rewrite fcont_comp_simpl. rewrite fcont_comp_simpl.
|
||||
refine (KLEISLIR_eq _ _ _).
|
||||
intros d0 d1 vd0 vd1. rewrite fcont_comp_simpl. rewrite PROD_map_simpl. simpl.
|
||||
rewrite vd0 in vd1. assert (vv:=vinj vd1). assert (stable f) as sf by auto. unfold stable in sf.
|
||||
apply (sf (h gg,d0) (h gg,d1)). auto.
|
||||
intros d0 vd0. exists d0. auto. intros aa vv. eexists ; apply vv.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_unit: forall D E F G (f:Dprod D E -C-> DL F) (g:G -C-> D) h,
|
||||
KLEISLIL f << <| eta << g, h |> == f << <| g, h |>.
|
||||
intros.
|
||||
unfold KLEISLIL. unfold RStrength.
|
||||
repeat (rewrite fcont_comp_assoc).
|
||||
refine (fcont_eq_intro _). intros gg. unfold PROD_map. repeat (rewrite fcont_comp_simpl). repeat (rewrite PROD_fun_simpl).
|
||||
rewrite uncurry_simpl. unfold Smash. repeat (rewrite fcont_comp_simpl).
|
||||
rewrite eta_val. rewrite FST_simpl, ID_simpl. simpl. rewrite SND_simpl, eta_val. simpl.
|
||||
refine (Oeq_trans (app_eq_compat (Oeq_refl _) (Operator2_simpl _ _ _)) _).
|
||||
rewrite eta_val. rewrite kleisli_simpl. rewrite kleisli_Val. auto.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIR_unit: forall D E F G (f:Dprod E D -C-> DL F) (g:G -C-> D) h,
|
||||
KLEISLIR f << PROD_fun h (eta << g) == f << PROD_fun h g.
|
||||
intros.
|
||||
unfold KLEISLIR. unfold LStrength.
|
||||
repeat (rewrite fcont_comp_assoc).
|
||||
refine (fcont_eq_intro _). intros gg. repeat (rewrite fcont_comp_simpl). repeat (rewrite PROD_map_simpl).
|
||||
rewrite uncurry_simpl. unfold Smash. rewrite ID_simpl. rewrite eta_val. simpl.
|
||||
refine (Oeq_trans (app_eq_compat (Oeq_refl _) (Operator2_simpl _ _ _)) _).
|
||||
rewrite eta_val. rewrite kleisli_simpl. rewrite kleisli_Val. auto.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIR_ValVal: forall D E F (f:Dprod D E -C-> DL F) e d r, KLEISLIR f (d,e) == Val r ->
|
||||
exists ee, e == Val ee /\ f (d,ee) == Val r.
|
||||
intros D E F f e d r. unfold KLEISLIR. unfold LStrength.
|
||||
repeat (rewrite fcont_comp_simpl). rewrite kleisli_simpl.
|
||||
intros klv. destruct (kleisliValVal klv) as [dd [Pdd PP]].
|
||||
rewrite PROD_map_simpl in Pdd. rewrite eta_val in Pdd.
|
||||
simpl in Pdd. rewrite ID_simpl in Pdd. simpl in Pdd.
|
||||
rewrite uncurry_simpl in Pdd. unfold Smash in Pdd.
|
||||
destruct (Operator2Val Pdd) as [d1 [e1 [Pd1 [Pe1 XX]]]]. exists e1.
|
||||
split. auto. rewrite eta_val in XX. assert (YY:=vinj XX).
|
||||
assert (stable f) as sf by auto. rewrite <- (vinj Pd1) in YY.
|
||||
specialize (sf _ _ YY).
|
||||
refine (Oeq_trans sf PP).
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_ValVal: forall D E F (f:Dprod D E -C-> DL F) e d r, KLEISLIL f (d,e) == Val r ->
|
||||
exists dd, d == Val dd /\ f (dd,e) == Val r.
|
||||
intros D E F f e d r. unfold KLEISLIL. unfold RStrength.
|
||||
repeat (rewrite fcont_comp_simpl). rewrite kleisli_simpl.
|
||||
intros klv. destruct (kleisliValVal klv) as [dd [Pdd PP]].
|
||||
rewrite PROD_map_simpl in Pdd. rewrite eta_val in Pdd.
|
||||
simpl in Pdd. rewrite ID_simpl in Pdd. simpl in Pdd.
|
||||
rewrite uncurry_simpl in Pdd. unfold Smash in Pdd.
|
||||
destruct (Operator2Val Pdd) as [d1 [e1 [Pd1 [Pe1 XX]]]]. exists d1.
|
||||
split. auto. rewrite eta_val in XX. assert (YY:=vinj XX).
|
||||
assert (stable f) as sf by auto. rewrite <- (vinj Pe1) in YY.
|
||||
specialize (sf _ _ YY).
|
||||
refine (Oeq_trans sf PP).
|
||||
Qed.
|
||||
|
||||
(*
|
||||
Definition fcont_alg : cpo -> alg -> alg.
|
||||
intros D A. exists (D -C-> A) (curry (strength A << (KLEISLIL (eta _ << (EV D A))))).
|
||||
apply fcont_eq_intro. intros f.
|
||||
rewrite fcont_comp_simpl. rewrite ID_simpl. simpl.
|
||||
apply fcont_eq_intro. intros d. rewrite curry_simpl. rewrite fcont_comp_simpl.
|
||||
assert (kl:=fcont_eq_elim (@KLEISLIL_unit _ _ _ _ (eta A << EV D A) (K _ f) (ID _)) d).
|
||||
repeat (rewrite fcont_comp_simpl in kl).
|
||||
assert ((PROD_fun (A:=DL (D -C-> A)) (B:=D) (C:=D)
|
||||
(eta (D -C-> A) << K D (D2:=D -C-> A) f)
|
||||
(ID D) d) == (eta (D -C-> A) f, d)) as F by auto.
|
||||
assert (X:=app_eq_compat (Oeq_refl (KLEISLIL (eta A << EV D A))) F).
|
||||
rewrite X in kl. clear F X.
|
||||
rewrite (app_eq_compat (Oeq_refl (strength A)) kl). clear kl.
|
||||
assert (X:=fcont_eq_elim (st_eta A) (EV D A
|
||||
(PROD_fun (A:=D -C-> A) (B:=D) (C:=D) (K D (D2:=D -C-> A) f)
|
||||
(ID D) d))). auto.
|
||||
|
||||
rewrite curry_compr.
|
||||
refine (Oeq_sym _).
|
||||
apply curry_unique.
|
||||
rewrite fcont_comp_assoc.
|
||||
|
||||
|
||||
|
||||
|
||||
apply fcont_eq_intro.
|
||||
intros f''. repeat (rewrite fcont_comp_simpl).
|
||||
refine (fcont_eq_intro _). intros d.
|
||||
repeat (rewrite curry_simpl).
|
||||
repeat (rewrite fcont_comp_simpl).
|
||||
assert (X:=st_kl A).
|
||||
|
||||
*)
|
||||
|
||||
Record fstricti (D1 D2 : cpo) (PD1:Pointed D1) (PD2:Pointed D2) : Type
|
||||
:= mk_fstricti {fstrictit : D1 -c> D2; fstrict : strict fstrictit}.
|
||||
|
||||
Definition fstricti_fun (D1 D2 : cpo) (PD1:Pointed D1) (PD2:Pointed D2) (f:fstricti PD1 PD2) : D1 -> D2 := fun x =>
|
||||
fstrictit f x.
|
||||
Coercion fstricti_fun : fstricti >-> Funclass.
|
||||
|
||||
Definition fstrict_ord (D1 D2:cpo) (PD1:Pointed D1) (PD2:Pointed D2) : ord.
|
||||
intros D1 D2 PD1 PD2.
|
||||
exists (fstricti PD1 PD2) (fun (f g:fstricti PD1 PD2) => fstrictit f <= fstrictit g).
|
||||
intros; auto.
|
||||
intros; auto.
|
||||
apply Ole_trans with (fstrictit y); auto.
|
||||
Defined.
|
||||
|
||||
Infix "-s>" := fstrict_ord (at level 30, right associativity) : O_scope.
|
||||
|
||||
(* Now Christine has a whole bunch of auxiliary lemmas for continuous
|
||||
maps, which I copy here, assuming they'll turn out to be useful
|
||||
*)
|
||||
|
||||
Lemma fstrict_le_intro : forall (D1 D2:cpo) (PD1:Pointed D1) (PD2:Pointed D2) (f g : PD1 -s> PD2), (forall x, f x <=
|
||||
g x) -> f <= g.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma fstrict_le_elim : forall (D1 D2:cpo) (PD1:Pointed D1) (PD2:Pointed D2) (f g : PD1 -s> PD2), f <= g -> forall x
|
||||
, f x <= g x.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma fstrict_eq_intro : forall (D1 D2:cpo) (PD1:Pointed D1) (PD2:Pointed D2) (f g : PD1 -s> PD2), (forall x, f x ==
|
||||
g x) -> f == g.
|
||||
intros; apply Ole_antisym; apply fstrict_le_intro; auto.
|
||||
Save.
|
||||
|
||||
Lemma fstrict_eq_elim : forall (D1 D2:cpo) (PD1:Pointed D1) (PD2:Pointed D2) (f g : PD1 -s> PD2), f == g -> forall x
|
||||
, f x == g x.
|
||||
intros; apply Ole_antisym; apply fstrict_le_elim; auto.
|
||||
Save.
|
||||
|
||||
Lemma fstrict_monotonic : forall (D1 D2:cpo) (PD1:Pointed D1) (PD2:Pointed D2) (f : PD1 -s> PD2) (x y : D1),
|
||||
x <= y -> f x <= f y.
|
||||
intros; apply (fmonotonic (fcontit (fstrictit f)) H).
|
||||
Save.
|
||||
Hint Resolve fstrict_monotonic.
|
||||
|
||||
Lemma fstrict_stable : forall (D1 D2:cpo) (PD1:Pointed D1) (PD2:Pointed D2) (f : PD1 -s> PD2) (x y : D1),
|
||||
x == y -> f x == f y.
|
||||
intros; apply (fmon_stable (fcontit (fstrictit f)) H).
|
||||
Save.
|
||||
Hint Resolve fstrict_monotonic.
|
||||
|
||||
Lemma strict_lub : forall (D1 D2 : cpo) (PD1:Pointed D1) (PD2:Pointed D2) (f:natO -m> (D1 -c> D2)),
|
||||
(forall n, strict (f n)) ->
|
||||
strict (lub (c:=D1-C->D2) f).
|
||||
intros. unfold strict.
|
||||
rewrite fcont_lub_simpl.
|
||||
split ; auto.
|
||||
apply lub_le.
|
||||
intro n.
|
||||
rewrite fcont_app_simpl.
|
||||
unfold strict in H.
|
||||
rewrite (H n).
|
||||
auto.
|
||||
Defined.
|
||||
|
||||
Definition Fstrictit (D1 D2 : cpo) (PD1:Pointed D1) (PD2:Pointed D2) : (PD1 -s> PD2)-m> (D1-c> D2).
|
||||
intros D1 D2 PD1 PD2; exists (fstrictit (D1:=D1) (D2:=D2) (PD1:=PD1) (PD2:=PD2)).
|
||||
auto.
|
||||
Defined.
|
||||
|
||||
Definition fstrict_lub (D1 D2 : cpo) (PD1:Pointed D1) (PD2:Pointed D2) : (natO -m> PD1 -s> PD2) -> (PD1 -s> PD2).
|
||||
intros D1 D2 PD1 PD2 f.
|
||||
exists (lub (c:= D1 -C->D2) (Fstrictit PD1 PD2 @ f)).
|
||||
apply strict_lub.
|
||||
intros; simpl.
|
||||
apply (fstrict (f n)).
|
||||
Defined.
|
||||
|
||||
Definition fstrict_cpo (D1 D2:cpo) (PD1:Pointed D1) (PD2:Pointed D2) : cpo.
|
||||
intros D1 D2 PD1 PD2; exists (fstrict_ord PD1 PD2)
|
||||
(fstrict_lub (D1:=D1) (D2:=D2) (PD1:=PD1) (PD2:=PD2));
|
||||
simpl;intros; auto.
|
||||
apply (le_lub (((Fcontit D1 D2 @ (Fstrictit PD1 PD2 @ c)) <_> x))
|
||||
(n)).
|
||||
Defined.
|
||||
|
||||
Infix "-S->" := fstrict_cpo (at level 30, right associativity) : O_scope.
|
||||
|
||||
Definition fstrict_comp : forall (D1 D2 D3:cpo) {PD1:Pointed D1} {PD2:Pointed D2} {PD3:Pointed D3},
|
||||
(PD2 -s> PD3) -> (PD1-s> PD2) -> PD1 -s> PD3.
|
||||
intros D1 D2 D3 PD1 DP2 PD3 f g.
|
||||
exists ((fstrictit f) << (fstrictit g)).
|
||||
red.
|
||||
rewrite fcont_comp_simpl.
|
||||
apply Oeq_trans with (fstrictit f (PBot)).
|
||||
apply fcont_stable.
|
||||
apply (fstrict g).
|
||||
apply (fstrict f).
|
||||
Defined.
|
||||
|
||||
Lemma strictVal: forall D E (f:DL D -c> DL E), (forall d e, f d == Val e -> exists dd, d == Val dd) -> strict f.
|
||||
intros. unfold strict.
|
||||
split ; auto.
|
||||
|
||||
simpl. apply DLless_cond.
|
||||
intros xx Fx. specialize (H _ _ Fx). destruct H as [d incon].
|
||||
simpl in incon.
|
||||
assert (yy := eqValpred incon).
|
||||
destruct yy as [n [dd [pp _]]].
|
||||
assert False as F. induction n.
|
||||
simpl in pp. rewrite DL_bot_eq in pp. inversion pp.
|
||||
rewrite DL_bot_eq in pp. simpl in pp. auto. inversion F.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_bot: forall D E (f:D -c> DL E), kleisli f (@PBot (DL D) _) == PBot.
|
||||
intros.
|
||||
refine (strictVal (f:=kleisli f) _).
|
||||
intros d e. rewrite kleisli_simpl.
|
||||
intros kl. destruct (kleisliValVal kl) as [dd [P _]]. exists dd. auto.
|
||||
Qed.
|
||||
|
979
papers/domains-2009/PredomLift.v
Executable file
979
papers/domains-2009/PredomLift.v
Executable file
|
@ -0,0 +1,979 @@
|
|||
(*==========================================================================
|
||||
Lifiting
|
||||
==========================================================================*)
|
||||
Require Import utility.
|
||||
Require Import PredomCore.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
(** ** Lifting *)
|
||||
|
||||
Section Lift_cpo.
|
||||
Variable D : cpo.
|
||||
|
||||
(** ** Definition of underlying stream type *)
|
||||
|
||||
(** printing Eps %\eps% *)
|
||||
(** printing Val %\val% *)
|
||||
(** printing Stream %\Stream% *)
|
||||
CoInductive Stream : Type := Eps : Stream -> Stream | Val : D -> Stream.
|
||||
|
||||
Lemma DL_inv : forall d, d = match d with Eps x => Eps x | Val d => Val d end.
|
||||
destruct d; auto.
|
||||
Save.
|
||||
Hint Resolve DL_inv.
|
||||
|
||||
(** ** Removing Eps steps *)
|
||||
|
||||
Definition pred d : Stream := match d with Eps x => x | Val _ => d end.
|
||||
|
||||
Fixpoint pred_nth (d:Stream) (n:nat) {struct n} : Stream :=
|
||||
match n with 0 => d
|
||||
|S m => match d with Eps x => pred_nth x m
|
||||
| Val _ => d
|
||||
end
|
||||
end.
|
||||
|
||||
Lemma pred_nth_val : forall x n, pred_nth (Val x) n = Val x.
|
||||
destruct n; simpl; trivial.
|
||||
Save.
|
||||
Hint Resolve pred_nth_val.
|
||||
|
||||
Lemma pred_nth_Sn_acc : forall n d, pred_nth d (S n) = pred_nth (pred d) n.
|
||||
destruct d; simpl; auto.
|
||||
Save.
|
||||
|
||||
Lemma pred_nth_Sn : forall n d, pred_nth d (S n) = pred (pred_nth d n).
|
||||
induction n; intros; auto.
|
||||
destruct d.
|
||||
exact (IHn d).
|
||||
rewrite pred_nth_val; rewrite pred_nth_val; simpl; trivial.
|
||||
Save.
|
||||
|
||||
(** ** Order *)
|
||||
CoInductive DLle : Stream -> Stream -> Prop :=
|
||||
| DLleEps : forall x y, DLle x y -> DLle (Eps x) (Eps y)
|
||||
| DLleEpsVal : forall x d, DLle x (Val d) -> DLle (Eps x) (Val d)
|
||||
| DLleVal : forall d d' n y, pred_nth y n = Val d' -> d <= d' -> DLle (Val d) y.
|
||||
|
||||
Hint Constructors DLle.
|
||||
|
||||
Lemma DLle_rec : forall R : Stream -> Stream -> Prop,
|
||||
(forall x y, R (Eps x) (Eps y) -> R x y) ->
|
||||
(forall x d, R (Eps x) (Val d) -> R x (Val d)) ->
|
||||
(forall d y, R (Val d) y -> exists n, exists d', pred_nth y n = Val d' /\ d <= d')
|
||||
-> forall x y, R x y -> DLle x y.
|
||||
intros R REps REpsVal RVal; cofix; destruct x; intros.
|
||||
destruct y; intros.
|
||||
apply DLleEps; apply DLle_rec; auto.
|
||||
apply DLleEpsVal; apply DLle_rec; auto.
|
||||
case (RVal t y H); intros n H1.
|
||||
destruct H1.
|
||||
apply DLleVal with x n; auto.
|
||||
intuition.
|
||||
intuition.
|
||||
Save.
|
||||
|
||||
(** *** Properties of the order *)
|
||||
|
||||
Lemma DLle_refl : forall x, DLle x x.
|
||||
intros x. apply DLle_rec with (R:= fun x y => forall n, pred_nth x n = pred_nth y n) ; try (clear x).
|
||||
intros x y C n. specialize (C (S n)). simpl in C. auto.
|
||||
intros x d C n. specialize (C (S n)). simpl in C. rewrite C. case n ; auto.
|
||||
intros d y C. exists 0. exists d. simpl. specialize (C 0). simpl in C. auto.
|
||||
intros n ; auto.
|
||||
Save.
|
||||
Hint Resolve DLle_refl.
|
||||
|
||||
Lemma DLvalval : forall d x, DLle (Val d) x ->
|
||||
exists n, exists d', pred_nth x n = Val d' /\ d<=d'.
|
||||
intros d x H.
|
||||
inversion H.
|
||||
exists n; exists d'.
|
||||
intuition.
|
||||
Qed.
|
||||
|
||||
Lemma pred_nth_epsS n x y: Eps y = pred_nth x n -> y = pred_nth x (S n).
|
||||
intros n. induction n. intros x y E. simpl in E. rewrite <- E. simpl. auto.
|
||||
intros x. case x ; clear x. intros x y. intros E. simpl in E. specialize (IHn _ _ E). auto.
|
||||
intros x y. simpl. intros incon. inversion incon.
|
||||
Qed.
|
||||
|
||||
Lemma pred_nth_valS n d y: Val d = pred_nth y n -> Val d = pred_nth y (S n).
|
||||
intros n ; induction n. intros d y E. simpl in E. rewrite <- E. auto.
|
||||
intros d y. case y ; clear y. intros y E. simpl in E. specialize (IHn _ _ E). rewrite IHn. auto.
|
||||
intros dd E. simpl in E. inversion E. rewrite <- H0 in *. clear dd H0 E. simpl. auto.
|
||||
Qed.
|
||||
|
||||
Lemma pred_nth_comp y : forall m0 n1 m1 : nat,
|
||||
m0 = n1 + m1 -> pred_nth y m0 = pred_nth (pred_nth y n1) m1.
|
||||
intros y m0. induction m0. intros n1 m1. simpl. intros E. assert (n1 = 0) by omega. rewrite H in *.
|
||||
simpl. assert (m1 = 0) as A by omega. rewrite A in *. auto.
|
||||
intros n1 m1. case m1 ; clear m1. case n1 ; clear n1. simpl. intros incon. inversion incon.
|
||||
intros m1 E. assert (m0 = m1) as EE by omega. rewrite EE in *. auto.
|
||||
intros m1. intros E. assert (m0 = n1 + m1) as A by omega. specialize (IHm0 _ _ A).
|
||||
clear E A. generalize IHm0. clear IHm0. case_eq (pred_nth y m0).
|
||||
intros d E. rewrite <- (pred_nth_epsS (sym_eq E)). intros EE. rewrite (pred_nth_epsS EE). auto.
|
||||
intros d E EE.
|
||||
rewrite <- (pred_nth_valS (sym_eq E)). rewrite (pred_nth_valS EE). auto.
|
||||
Qed.
|
||||
|
||||
Lemma DLle_pred_nth x y n : DLle x y -> DLle (pred_nth x n) (pred_nth y n).
|
||||
intros x y n. induction n. simpl. auto. intros C.
|
||||
specialize (IHn C). inversion IHn.
|
||||
rewrite (pred_nth_epsS H) in *. clear H x0. rewrite (pred_nth_epsS H0) in *. clear y0 H0. auto.
|
||||
rewrite (pred_nth_epsS H) in *. clear H x0. rewrite H0 in *.
|
||||
rewrite <- (pred_nth_valS H0). rewrite <- H0 in *. auto.
|
||||
rewrite <- H2 in *. rewrite <- H in *.
|
||||
rewrite <- (pred_nth_valS H).
|
||||
apply DLleVal with (n:=n0) (d':=d').
|
||||
rewrite H2 in H0. rewrite <- (@pred_nth_comp y (n+n0)) in H0 ; auto.
|
||||
rewrite <- (@pred_nth_comp y (S n+n0)) ; auto. rewrite (pred_nth_valS (sym_eq H0)). auto. auto.
|
||||
Qed.
|
||||
|
||||
Lemma DLleEps_right : forall x y, DLle x y -> DLle x (Eps y).
|
||||
intros x y C. apply DLle_rec with (R:= fun x y => forall n, DLle (pred_nth x n) (pred_nth y (S n))).
|
||||
clear x y C. intros x y C n. specialize (C (S n)). auto.
|
||||
clear x y C. intros x y C n. specialize (C (S n)). auto.
|
||||
clear x y C. intros d y C. specialize (C 0).
|
||||
assert (DLle (Val d) (pred_nth y 1)) as CC by auto. clear C.
|
||||
destruct (DLvalval CC) as [n [dd [X Y]]]. exists (S n). exists dd. split ; auto.
|
||||
generalize X. clear X. destruct y. simpl. auto. simpl. case n ; auto.
|
||||
intros n. simpl.
|
||||
apply DLle_pred_nth ; auto.
|
||||
Qed.
|
||||
Hint Resolve DLleEps_right.
|
||||
|
||||
Lemma DLleEps_left : forall x y, DLle x y -> DLle (Eps x) y.
|
||||
intros x y C. apply DLle_rec with (R:= fun x y => forall n, DLle (pred_nth x (S n)) (pred_nth y n)).
|
||||
clear x y C. intros x y C n. specialize (C (S n)). auto.
|
||||
clear x y C. intros x y C n. specialize (C (S n)). destruct n ; auto.
|
||||
clear x y C. intros d y C. specialize (C 0).
|
||||
simpl in C.
|
||||
destruct (DLvalval C) as [n [dd [X Y]]]. exists n. exists dd. split ; auto.
|
||||
intros n. simpl. apply DLle_pred_nth ; auto.
|
||||
Qed.
|
||||
Hint Resolve DLleEps_left.
|
||||
|
||||
Lemma DLle_pred_left : forall x y, DLle x y -> DLle (pred x) y.
|
||||
destruct x; destruct y; simpl; intros; trivial.
|
||||
inversion H; auto.
|
||||
inversion H; trivial.
|
||||
Save.
|
||||
|
||||
Lemma DLle_pred_right : forall x y, DLle x y -> DLle x (pred y).
|
||||
destruct x; destruct y; simpl; intros; trivial.
|
||||
inversion H; auto.
|
||||
inversion H; trivial.
|
||||
destruct n; simpl in H1.
|
||||
discriminate H1.
|
||||
apply DLleVal with d' n; auto.
|
||||
Save.
|
||||
|
||||
Hint Resolve DLle_pred_left DLle_pred_right.
|
||||
|
||||
Lemma DLle_pred : forall x y, DLle x y -> DLle (pred x) (pred y).
|
||||
auto.
|
||||
Save.
|
||||
|
||||
Hint Resolve DLle_pred.
|
||||
|
||||
Lemma DLle_pred_nth_left : forall n x y, DLle x y -> DLle (pred_nth x n) y.
|
||||
induction n; intros.
|
||||
simpl; auto.
|
||||
rewrite pred_nth_Sn; auto.
|
||||
Save.
|
||||
|
||||
Lemma DLle_pred_nth_right : forall n x y,
|
||||
DLle x y -> DLle x (pred_nth y n).
|
||||
induction n; intros.
|
||||
simpl; auto.
|
||||
rewrite pred_nth_Sn; auto.
|
||||
Save.
|
||||
|
||||
Hint Resolve DLle_pred_nth_left DLle_pred_nth_right.
|
||||
|
||||
|
||||
(* should be called leq *)
|
||||
Lemma DLleVal_leq : forall x y, DLle (Val x) (Val y) -> x <= y.
|
||||
intros x y H; inversion H.
|
||||
destruct n; simpl in H1; injection H1;intro; subst y; auto.
|
||||
Save.
|
||||
Hint Immediate DLleVal_leq.
|
||||
|
||||
(* New *)
|
||||
Lemma DLle_leVal : forall x y, x<=y -> DLle (Val x) (Val y).
|
||||
intros x y H.
|
||||
apply (DLle_rec (R := fun x' y' => x'=Val x /\ y'=Val y)).
|
||||
intros.
|
||||
destruct H0; discriminate.
|
||||
intros.
|
||||
destruct H0; discriminate.
|
||||
intros.
|
||||
destruct H0.
|
||||
subst y0.
|
||||
exists 0.
|
||||
exists y.
|
||||
simpl.
|
||||
split; trivial.
|
||||
injection H0.
|
||||
intro; subst d; auto.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Require Import utility.
|
||||
|
||||
Lemma DLnvalval : forall n y d z, pred_nth y n = Val d -> DLle y z ->
|
||||
exists m, exists d', pred_nth z m = Val d' /\ d <= d'.
|
||||
induction n.
|
||||
intros.
|
||||
dcase y.
|
||||
intros.
|
||||
subst y.
|
||||
simpl in H.
|
||||
discriminate.
|
||||
intros.
|
||||
subst y.
|
||||
simpl in H.
|
||||
rewrite H in H0.
|
||||
apply (DLvalval H0).
|
||||
|
||||
intros.
|
||||
apply (IHn (pred y) d z).
|
||||
rewrite <- (pred_nth_Sn_acc).
|
||||
assumption.
|
||||
apply DLle_pred_left.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
Lemma DLleEpsn n x z xx : pred_nth x n = Val xx -> DLle (Val xx) z -> DLle x z.
|
||||
intros n. induction n.
|
||||
intros x z xx. simpl. intros C. rewrite C. auto.
|
||||
intros x. case x ; clear x. intros d z xx. simpl. intros C DD.
|
||||
specialize (IHn _ z _ C DD).
|
||||
generalize IHn. case z. intros zz CC. apply DLleEps. apply (DLle_pred_right CC).
|
||||
intros zz CC. apply DLleEpsVal. apply CC.
|
||||
intros x z xx C. simpl in C. inversion C. auto.
|
||||
Qed.
|
||||
|
||||
Lemma DLle_trans : forall x y z, DLle x y -> DLle y z -> DLle x z.
|
||||
intros x y z D0 D1.
|
||||
apply DLle_rec with (R:=fun x z => (forall xx n, pred_nth x n = Val xx -> DLle x z)).
|
||||
clear x y z D0 D1. intros x y C xx n X.
|
||||
specialize (C xx (S n) X).
|
||||
apply (DLle_pred C).
|
||||
|
||||
clear x y z D0 D1. intros x y C xx n X.
|
||||
specialize (C _ (S n) X).
|
||||
apply (DLle_pred_left C).
|
||||
|
||||
intros d yy C.
|
||||
specialize (C d 0).
|
||||
destruct (fun Z => DLvalval (C Z)) as [n [dd [ZZ Z]]] ; auto.
|
||||
exists n. exists dd. split ; auto.
|
||||
|
||||
intros xx n C.
|
||||
destruct (DLnvalval C D0) as [m [yy [P Q]]].
|
||||
destruct (DLnvalval P D1) as [k [zz [Pz Qz]]].
|
||||
apply (DLleEpsn C). apply (DLleVal Pz). apply (Ole_trans Q Qz).
|
||||
Qed.
|
||||
|
||||
(** *** Declaration of the ordered set *)
|
||||
Definition DL_ord : ord.
|
||||
exists Stream DLle; intros; auto.
|
||||
apply DLle_trans with y; trivial.
|
||||
Defined.
|
||||
|
||||
Canonical Structure DL_ord.
|
||||
|
||||
|
||||
(** ** Definition of the cpo structure *)
|
||||
|
||||
Lemma eq_Eps : forall x, x == Eps x.
|
||||
intros; apply Ole_antisym; repeat red; auto.
|
||||
Save.
|
||||
Hint Resolve eq_Eps.
|
||||
|
||||
(** *** Bottom is given by an infinite chain of Eps *)
|
||||
(** printing DL_bot %\DLbot% *)
|
||||
CoFixpoint DL_bot : DL_ord := Eps DL_bot.
|
||||
|
||||
Lemma DL_bot_eq : DL_bot = Eps DL_bot.
|
||||
transitivity match DL_bot with Eps x => Eps x | Val d => Val d end; auto.
|
||||
Save.
|
||||
|
||||
Lemma DLless_cond : forall (x y : DL_ord), (forall xx, x == Val xx -> x <= y) -> DLle x y.
|
||||
intros x y P. apply DLle_rec with (R:=fun x y => forall xx, x == Val xx -> x <= y).
|
||||
intros d0 d1 IH. intros d00 dv.
|
||||
rewrite (eq_Eps d0) in dv. rewrite (eq_Eps d0).
|
||||
specialize (IH _ dv).
|
||||
rewrite (eq_Eps d1). auto.
|
||||
|
||||
intros d0 d1 IH dd dv. rewrite (eq_Eps d0) in dv.
|
||||
specialize (IH _ dv). rewrite (eq_Eps d0). auto.
|
||||
|
||||
intros d0 d1 IH.
|
||||
specialize (IH _ (Oeq_refl _)).
|
||||
auto using (DLvalval IH).
|
||||
apply P.
|
||||
Qed.
|
||||
|
||||
Lemma DL_bot_least : forall x, DL_bot <= x.
|
||||
intros x. apply DLless_cond.
|
||||
intros xx [_ C].
|
||||
destruct (DLvalval C) as [n [dd [P Q]]].
|
||||
assert False as F. induction n. simpl in P. rewrite DL_bot_eq in P. inversion P.
|
||||
rewrite DL_bot_eq in P. simpl in P. apply (IHn P). inversion F.
|
||||
Qed.
|
||||
|
||||
(** *** More properties of elements in the lifted domain *)
|
||||
|
||||
(* This is essentially the same, when fixed up, as DLvalval that I proved above
|
||||
I've repeated it here, though
|
||||
*)
|
||||
Lemma DLle_Val_exists_pred :
|
||||
forall x d, Val d <= x -> exists k, exists d', pred_nth x k = Val d'
|
||||
/\ d <= d'.
|
||||
intros x d H; inversion H; eauto.
|
||||
Save.
|
||||
|
||||
Lemma Val_exists_pred_le :
|
||||
forall x d, (exists k, pred_nth x k = Val d) -> Val d <= x.
|
||||
destruct 1; intros.
|
||||
apply DLleVal with d x0; trivial.
|
||||
Save.
|
||||
Hint Immediate DLle_Val_exists_pred Val_exists_pred_le.
|
||||
|
||||
Lemma Val_exists_pred_eq :
|
||||
forall x d, (exists k, pred_nth x k = Val d) -> (Val d:DL_ord) == x.
|
||||
intros.
|
||||
split.
|
||||
auto.
|
||||
destruct H.
|
||||
rewrite <- H.
|
||||
apply DLle_pred_nth_right.
|
||||
auto.
|
||||
Save.
|
||||
|
||||
(** *** Construction of least upper bounds *)
|
||||
|
||||
Definition isEps x := match x with Eps _ => True | _ => False end.
|
||||
|
||||
Lemma isEps_Eps : forall x, isEps (Eps x).
|
||||
repeat red; auto.
|
||||
Save.
|
||||
|
||||
Lemma not_isEpsVal : forall d, ~ (isEps (Val d)).
|
||||
repeat red; auto.
|
||||
Save.
|
||||
Hint Resolve isEps_Eps not_isEpsVal.
|
||||
|
||||
Lemma isEps_dec : forall x, {d:D|x=Val d}+{isEps x}.
|
||||
destruct x; auto.
|
||||
left.
|
||||
exists t; auto.
|
||||
Defined.
|
||||
|
||||
(* Slightly more radical rewrite, trying to use equality *)
|
||||
|
||||
Lemma allvalsfromhere : forall (c:natO -m> DL_ord) n d i,
|
||||
c n == Val d -> exists d', c (n+i) == Val d' /\ d <= d'.
|
||||
intros c n d i (Hnv1,Hnv2).
|
||||
assert (((Val d):DL_ord) <= c (n+i)).
|
||||
apply Ole_trans with (c n); auto.
|
||||
apply fmonotonic.
|
||||
intuition.
|
||||
destruct (DLle_Val_exists_pred H) as [k [d' Hkd']].
|
||||
exists d'.
|
||||
split.
|
||||
apply Oeq_sym; apply Val_exists_pred_eq; firstorder.
|
||||
tauto.
|
||||
Qed.
|
||||
|
||||
Definition hasVal x := exists j, exists d, pred_nth x j = Val d.
|
||||
Definition valuable := {x | hasVal x}.
|
||||
|
||||
Definition projj : valuable -> DL_ord.
|
||||
intro.
|
||||
destruct X.
|
||||
exact x.
|
||||
Defined.
|
||||
|
||||
|
||||
(* Redoing extract using constructive epsilon *)
|
||||
Require Import ConstructiveEpsilon.
|
||||
|
||||
|
||||
Definition findindexandval (x : valuable) : {kd | match kd with (k,d) => (pred_nth (projj x) k) = Val d end}.
|
||||
intro x; destruct x.
|
||||
unfold hasVal in h; simpl.
|
||||
cut {n | exists d, pred_nth x n = Val d}.
|
||||
intro; destruct H.
|
||||
dcase (pred_nth x x0).
|
||||
intros. rewrite H in e. absurd (exists d, Eps s = Val d).
|
||||
intro.
|
||||
destruct H0.
|
||||
discriminate.
|
||||
assumption.
|
||||
intros.
|
||||
exists (x0,t).
|
||||
assumption.
|
||||
|
||||
apply constructive_indefinite_description_nat.
|
||||
intro.
|
||||
case (isEps_dec (pred_nth x x0)).
|
||||
intro; left.
|
||||
destruct s.
|
||||
exists x1.
|
||||
assumption.
|
||||
intro H; right.
|
||||
intro; destruct H0.
|
||||
rewrite H0 in H; simpl in H.
|
||||
contradiction.
|
||||
exact h.
|
||||
Defined.
|
||||
|
||||
|
||||
Definition extract (x : valuable) : D.
|
||||
intro x; generalize (findindexandval x); intro.
|
||||
destruct X.
|
||||
destruct x0.
|
||||
exact t.
|
||||
Defined.
|
||||
|
||||
|
||||
Lemma extractworks : forall x, projj x == Val (extract x).
|
||||
intro x; unfold extract.
|
||||
dcase (findindexandval x).
|
||||
intros.
|
||||
destruct x0.
|
||||
apply Oeq_sym.
|
||||
apply Val_exists_pred_eq.
|
||||
exists n.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma vinj : forall d d', Val d == Val d' -> d == d'.
|
||||
intros d d' (H1,H2).
|
||||
split; auto.
|
||||
Qed.
|
||||
|
||||
Lemma vleinj : forall d d', Val d <= Val d' -> d <= d'.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma extractmono : forall (x y : valuable), (projj x) <= (projj y) -> extract x <= extract y.
|
||||
intros x y H.
|
||||
apply vleinj.
|
||||
rewrite <- (extractworks x).
|
||||
rewrite <- (extractworks y).
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
(* This is the simpler one that just takes an equality *)
|
||||
|
||||
Lemma eqValpred : forall x d, x == Val d -> exists k, exists d', pred_nth x k = Val d' /\ d==d'.
|
||||
intros.
|
||||
destruct H as (H1,H2).
|
||||
destruct (DLle_Val_exists_pred H2) as [k [d' HH]].
|
||||
exists k; exists d';split.
|
||||
tauto.
|
||||
split.
|
||||
tauto.
|
||||
apply vleinj.
|
||||
destruct HH.
|
||||
rewrite <- H.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Definition makechain (c:natO -m> DL_ord) k d (Hck:c k == Val d) : nat -> D.
|
||||
intros c k d Hck i.
|
||||
eapply extract.
|
||||
exists (c (k+i)).
|
||||
destruct (allvalsfromhere i Hck) as [d' [Hd' HHH]].
|
||||
unfold hasVal.
|
||||
generalize (eqValpred Hd').
|
||||
firstorder.
|
||||
Defined.
|
||||
|
||||
Definition makechainm (c:natO -m> DL_ord) k d (Hck:c k == Val d) : natO -m> D.
|
||||
intros.
|
||||
exists (makechain Hck).
|
||||
unfold monotonic.
|
||||
intros.
|
||||
unfold makechain.
|
||||
apply extractmono.
|
||||
simpl.
|
||||
apply (fmonotonic c).
|
||||
intuition.
|
||||
Defined.
|
||||
|
||||
|
||||
Definition cpred (c:natO -m> DL_ord) : natO-m>DL_ord.
|
||||
intro c; exists (fun n => pred (c n)); red.
|
||||
intros x y H; apply DLle_pred.
|
||||
apply (fmonotonic c); auto.
|
||||
Defined.
|
||||
|
||||
Lemma fValIdx : forall (c:natO -m> DL_ord) (n : nat),
|
||||
{dk: (D*nat)%type | match dk with (d,k) => k<n /\ c k == Val d end} + {forall k, k<n -> isEps (c k)}.
|
||||
induction n.
|
||||
right; intros. destruct (lt_n_O k H).
|
||||
case IHn; intros.
|
||||
destruct s as [(d,k) [H ck]].
|
||||
left; exists (d,k).
|
||||
auto with arith.
|
||||
|
||||
case (isEps_dec (c n)); intros.
|
||||
destruct s as [d H]. left.
|
||||
exists (d,n); intuition.
|
||||
|
||||
right; intros.
|
||||
assert (L:=lt_n_Sm_le _ _ H). destruct L ; auto with arith.
|
||||
Defined.
|
||||
|
||||
Definition createchain (c:natO -m> DL_ord) (n:nat)
|
||||
(X:{dk: (D*nat)%type | match dk with (d,k) => k<n /\ c k == Val d end}) :
|
||||
(natO -m> D) :=
|
||||
match X with | exist (d,k) (conj Hk Hck) => @makechainm c k d Hck
|
||||
end.
|
||||
|
||||
CoFixpoint DL_lubn (c:natO-m> DL_ord) (n:nat) : DL_ord :=
|
||||
match fValIdx c n with inleft X => Val (lub (createchain X))
|
||||
| inright _ => Eps (DL_lubn (cpred c) (S n))
|
||||
end.
|
||||
|
||||
Lemma DL_lubn_inv : forall (c:natO-m> DL_ord) (n:nat), DL_lubn c n =
|
||||
match fValIdx c n with inleft X => Val (lub (createchain X))
|
||||
| inright _ => Eps (DL_lubn (cpred c) (S n))
|
||||
end.
|
||||
intros; rewrite (DL_inv (DL_lubn c n)).
|
||||
simpl; case (fValIdx c n); trivial.
|
||||
Qed.
|
||||
|
||||
Lemma makechainworks : forall (c:natO -m> DL_ord) k dk (H2:c k == Val dk) i (d:D), c (k+i) == (Val d) -> makechain H2 i == d.
|
||||
intros.
|
||||
apply vinj.
|
||||
rewrite <- H.
|
||||
apply Oeq_sym.
|
||||
unfold makechain.
|
||||
rewrite <- extractworks.
|
||||
simpl.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma chainluble : forall (c:natO -m> DL_ord) k dk (Hkdk : c k == Val dk) k' dk' (Hkdk': c k' == Val dk'),
|
||||
lub (makechainm Hkdk) <= lub (makechainm Hkdk').
|
||||
intros;apply lub_le.
|
||||
intro n.
|
||||
assert (Hkn := allvalsfromhere n Hkdk).
|
||||
destruct Hkn as (dkn,(Hckn,Hdkdkn)).
|
||||
setoid_replace (makechainm Hkdk n) with (dkn).
|
||||
destruct (allvalsfromhere (n+k) Hkdk') as [dkk'n [Hckk'n Hddd]].
|
||||
apply Ole_trans with dkk'n.
|
||||
apply vleinj; rewrite <- Hckn; rewrite <- Hckk'n.
|
||||
apply fmonotonic.
|
||||
intuition.
|
||||
setoid_replace dkk'n with (makechainm Hkdk' (n+k)).
|
||||
apply le_lub.
|
||||
unfold makechainm.
|
||||
apply Oeq_sym; apply makechainworks.
|
||||
assumption.
|
||||
unfold makechainm; apply makechainworks.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
Lemma chainlubinv : forall (c:natO -m> DL_ord) k dk (Hkdk : c k == Val dk) k' dk' (Hkdk': c k' == Val dk'),
|
||||
lub (makechainm Hkdk) == lub (makechainm Hkdk').
|
||||
intros;split;apply chainluble.
|
||||
Qed.
|
||||
|
||||
Lemma pred_lubn_Val : forall (d:D)(n k p:nat) (c:natO-m> DL_ord),
|
||||
(n <k+p)%nat -> pred_nth (c n) k = Val d
|
||||
-> exists d', DL_lubn c p == Val d'.
|
||||
|
||||
induction k; intros.
|
||||
(* Base case *)
|
||||
rewrite (DL_lubn_inv c p); simpl.
|
||||
simpl in H0.
|
||||
case (fValIdx c p); intros.
|
||||
case s; intros (d',k) (H1,H2); auto.
|
||||
simpl.
|
||||
exists (lub (makechainm H2)).
|
||||
auto.
|
||||
|
||||
absurd (c n = Val d).
|
||||
intro Hsilly.
|
||||
assert (isEps (c n)).
|
||||
apply i.
|
||||
omega.
|
||||
rewrite Hsilly in H1.
|
||||
simpl in H1.
|
||||
tauto.
|
||||
assumption.
|
||||
|
||||
(* Induction *)
|
||||
rewrite (DL_lubn_inv c p).
|
||||
case (fValIdx c p); intros.
|
||||
unfold createchain.
|
||||
dependent inversion s.
|
||||
simpl.
|
||||
destruct x.
|
||||
dependent inversion y.
|
||||
exists (lub (makechainm o)).
|
||||
auto.
|
||||
|
||||
assert (n<k + S p).
|
||||
omega.
|
||||
assert (pred_nth (cpred c n) k = Val d).
|
||||
unfold cpred.
|
||||
simpl.
|
||||
rewrite <- pred_nth_Sn_acc.
|
||||
assumption.
|
||||
specialize (IHk (S p) (cpred c) H1 H2 ).
|
||||
destruct IHk as (d',IH); exists d'.
|
||||
rewrite <-IH; auto.
|
||||
Qed.
|
||||
|
||||
Lemma eqDLeq : forall d d', d == d' -> Val d == (Val d').
|
||||
intros.
|
||||
destruct H; split; apply DLle_leVal.
|
||||
assumption.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
Lemma predeq : forall x, x == pred x.
|
||||
intro; split; auto.
|
||||
Qed.
|
||||
|
||||
Lemma cpredsamelub : forall (c:natO -m> DL_ord) k dk (H1:c k == Val dk) (H2: (cpred c) k == Val dk),
|
||||
lub (makechainm H1) == lub (makechainm H2).
|
||||
intros; split.
|
||||
apply lub_le; intro.
|
||||
setoid_replace (makechainm H1 n) with (makechainm H2 n).
|
||||
apply le_lub.
|
||||
assert (Hkn := allvalsfromhere n H1).
|
||||
destruct Hkn as (dkn,(Hckn,Hdkdkn)).
|
||||
unfold makechainm; simpl.
|
||||
setoid_rewrite (makechainworks H1 Hckn).
|
||||
assert (HH : ((cpred c) (k+n) == Val dkn)).
|
||||
unfold cpred; simpl.
|
||||
rewrite <- Hckn.
|
||||
apply Oeq_sym; apply predeq.
|
||||
setoid_rewrite (makechainworks H2 HH).
|
||||
auto.
|
||||
|
||||
apply lub_le; intro.
|
||||
setoid_replace (makechainm H2 n) with (makechainm H1 n).
|
||||
apply le_lub.
|
||||
destruct (allvalsfromhere n H2) as (dkn,(Hckn,Hdkdkn)).
|
||||
unfold makechainm; simpl.
|
||||
setoid_rewrite (makechainworks H2 Hckn).
|
||||
assert (HH : ((c) (k+n) == Val dkn)).
|
||||
unfold cpred in Hckn; simpl in Hckn.
|
||||
rewrite <- Hckn; apply predeq.
|
||||
setoid_rewrite (makechainworks H1 HH); auto.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma attempt2 : forall kl (c:natO -m> DL_ord) p d, pred_nth (DL_lubn c p) kl = Val d
|
||||
-> exists k, exists dk, exists Hkdk:c k == Val dk, d == lub (makechainm Hkdk).
|
||||
induction kl; intros.
|
||||
rewrite DL_lubn_inv in H; simpl in H.
|
||||
dcase (fValIdx c p).
|
||||
intros.
|
||||
rewrite H0 in H.
|
||||
destruct s; destruct x; inversion y.
|
||||
|
||||
unfold createchain in H.
|
||||
exists n; exists t.
|
||||
exists H2.
|
||||
clear H0.
|
||||
generalize H.
|
||||
clear H.
|
||||
dependent inversion y.
|
||||
intro.
|
||||
apply vinj.
|
||||
rewrite <- H.
|
||||
|
||||
apply eqDLeq.
|
||||
apply chainlubinv.
|
||||
|
||||
(* then we get an impossible case *)
|
||||
|
||||
Focus 2.
|
||||
(* this is the tricky one *)
|
||||
assert (HH := IHkl (cpred c) (S p) d).
|
||||
dcase (fValIdx c p).
|
||||
intros.
|
||||
|
||||
(* making it up as I go along *)
|
||||
rewrite DL_lubn_inv in H.
|
||||
rewrite H0 in H; simpl in H.
|
||||
clear H0.
|
||||
destruct s.
|
||||
destruct x.
|
||||
generalize H; dependent inversion y; intros.
|
||||
unfold createchain in H0.
|
||||
exists n; exists t; exists o.
|
||||
apply vinj.
|
||||
rewrite <- H0.
|
||||
auto.
|
||||
|
||||
clear IHkl.
|
||||
intros.
|
||||
rewrite DL_lubn_inv in H.
|
||||
rewrite H0 in H; simpl in H.
|
||||
specialize (HH H).
|
||||
(* AHA now we're getting somewhere at last, albeit rather clunkily! *)
|
||||
destruct HH as (k,(dk,(Hkdk,HH))).
|
||||
exists k; exists dk.
|
||||
|
||||
unfold cpred in Hkdk; simpl in Hkdk.
|
||||
assert (HHH : c k == Val dk).
|
||||
rewrite <- Hkdk.
|
||||
apply predeq.
|
||||
exists HHH.
|
||||
rewrite HH.
|
||||
|
||||
apply Oeq_sym.
|
||||
apply cpredsamelub.
|
||||
|
||||
(* now back to trivial one *)
|
||||
intros.
|
||||
rewrite H0 in H.
|
||||
discriminate.
|
||||
Qed.
|
||||
|
||||
Lemma pred_nth_sum : forall x m n, pred_nth x (m+n) = pred_nth (pred_nth x m) n.
|
||||
induction m.
|
||||
simpl.
|
||||
intuition.
|
||||
intro n; replace (S m + n) with (m + S n); intuition.
|
||||
rewrite IHm.
|
||||
rewrite (@pred_nth_Sn_acc n (pred_nth x m)).
|
||||
rewrite pred_nth_Sn.
|
||||
reflexivity.
|
||||
Qed.
|
||||
|
||||
(* first just wrap pred_lubn_Val to use equality *)
|
||||
Lemma chainVallubnVal : forall (c:natO -m> DL_ord) n d p, c n == Val d -> exists d', DL_lubn c p == Val d'.
|
||||
intros.
|
||||
destruct (eqValpred H) as (k,(d'',(Hcnk,Hdd''))).
|
||||
apply (pred_lubn_Val (n:=n) (k:=k+n+1) (p:=p) (c:=c) (d:=d'')).
|
||||
intuition.
|
||||
replace (k+n+1) with (k+(n+1)).
|
||||
rewrite pred_nth_sum.
|
||||
rewrite Hcnk.
|
||||
auto.
|
||||
omega.
|
||||
Qed.
|
||||
|
||||
Lemma chainVallubnlub : forall (c:natO -m> DL_ord) n d (H : c n == Val d) p, exists d', DL_lubn c p == Val d' /\ d' == lub (makechainm H).
|
||||
intros.
|
||||
destruct (chainVallubnVal p H) as (d',HH).
|
||||
exists d'.
|
||||
split. assumption.
|
||||
destruct (eqValpred HH) as (k,(d'',(Hk,Hd'd''))).
|
||||
destruct (attempt2 Hk) as (p',(dp',(Hp',Heq))).
|
||||
rewrite Hd'd''.
|
||||
rewrite Heq.
|
||||
apply chainlubinv.
|
||||
Qed.
|
||||
|
||||
|
||||
Definition DL_lub (c:natO-m> DL_ord) := DL_lubn c 1.
|
||||
|
||||
|
||||
Lemma DL_lub_upper : forall c:natO-m> DL_ord, forall n, c n <= DL_lub c.
|
||||
intros c n. apply DLless_cond.
|
||||
intros d C.
|
||||
destruct (chainVallubnlub C 1) as [d' [Ln L]].
|
||||
unfold DL_lub. rewrite Ln.
|
||||
assert (X:=makechainworks C). specialize (X 0 d). assert (n+0 = n) as A by omega. rewrite A in X. clear A.
|
||||
specialize (X C). apply Ole_trans with (y:=Val d). auto.
|
||||
assert (pred_nth (Val d') 0 = Val d') as Y by auto.
|
||||
apply (DLleVal Y). rewrite L.
|
||||
apply Ole_trans with (y:=(makechainm (c:=c) (k:=n) (d:=d) C) 0). simpl. auto.
|
||||
apply le_lub.
|
||||
Qed.
|
||||
|
||||
(* Just realized I should have had Val as a morphism ages ago *)
|
||||
Add Parametric Morphism : Val
|
||||
with signature (@Oeq D) ==> (@Oeq DL_ord)
|
||||
as Val_eq_compat.
|
||||
intros.
|
||||
apply eqDLeq.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism : Val
|
||||
with signature (@Ole D) ++> (@Ole DL_ord)
|
||||
as Val_le_compat.
|
||||
intros.
|
||||
apply DLle_leVal.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma DLle_Val_exists_eq : forall y d, Val d <= y -> exists d', y == Val d' /\ d <= d'.
|
||||
intros y d H; inversion H.
|
||||
exists d'.
|
||||
split.
|
||||
symmetry.
|
||||
apply Val_exists_pred_eq.
|
||||
exists n; trivial.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma DL_lub_least : forall (c : natO -m> DL_ord) a,
|
||||
(forall n, c n <= a) -> DL_lub c <= a.
|
||||
intros c a C.
|
||||
apply DLless_cond. intros d X.
|
||||
destruct (eqValpred X) as [n [dd [P Q]]].
|
||||
destruct (attempt2 P) as [m [d' [cm Y]]].
|
||||
apply Ole_trans with (y:= (Val dd)).
|
||||
refine (proj2 (Val_exists_pred_eq _)). exists n. auto.
|
||||
assert (Z:= C m). rewrite cm in Z.
|
||||
destruct (DLvalval Z) as [mm [e [R S]]].
|
||||
apply (DLleVal R). rewrite Y.
|
||||
apply lub_le.
|
||||
intros nn.
|
||||
assert (A:=makechainworks cm).
|
||||
specialize (A nn).
|
||||
clear S Y Z.
|
||||
destruct (allvalsfromhere nn cm) as [nnv [S T]].
|
||||
specialize (A _ S). rewrite A. clear A.
|
||||
assert (a == Val e) as E.
|
||||
apply Oeq_sym.
|
||||
apply (Val_exists_pred_eq). exists mm. apply R.
|
||||
assert (Z := C (m+nn)).
|
||||
rewrite E in Z. rewrite S in Z. apply (vleinj Z).
|
||||
Qed.
|
||||
|
||||
|
||||
(** *** Declaration of the lifted cpo *)
|
||||
|
||||
Definition DL : cpo.
|
||||
exists DL_ord DL_lub; intros.
|
||||
(* apply DL_bot_least. *)
|
||||
apply DL_lub_upper.
|
||||
apply DL_lub_least; trivial.
|
||||
Defined.
|
||||
|
||||
Lemma lubval: forall (c:natO -m> DL) d, lub c == Val d ->
|
||||
exists k, exists d', c k == Val d' /\ d' <= d.
|
||||
intros c d l.
|
||||
destruct l as [lubval1 lubval2].
|
||||
destruct (DLle_Val_exists_pred lubval2) as [k xx].
|
||||
destruct xx as [f' [lubval lf]].
|
||||
assert (lub c == Val f') as lubval'.
|
||||
apply Ole_antisym.
|
||||
eapply Ole_trans. apply lubval1. auto using DLle_leVal.
|
||||
eapply (DLleVal). apply lubval. apply Ole_refl.
|
||||
assert (f' <= d).
|
||||
destruct lubval' as [v1 v2].
|
||||
assert (Val f' <= Val d).
|
||||
apply (Ole_trans v2 lubval1).
|
||||
apply (DLleVal_leq H).
|
||||
assert (d == f') as feq. auto.
|
||||
clear H lubval1 lubval2 lf.
|
||||
|
||||
assert (xx:=attempt2 lubval).
|
||||
destruct xx as [n [f0 [chnval lf']]].
|
||||
exists n. exists f0. split. apply chnval.
|
||||
assert (w:=makechainworks chnval).
|
||||
specialize (w 0 f0). assert (n+0 = n) by omega. rewrite H in w.
|
||||
specialize (w chnval).
|
||||
destruct w as [_ L2].
|
||||
refine (Ole_trans L2 (Ole_trans (le_lub (makechainm chnval) 0) _)).
|
||||
rewrite <- lf'. auto.
|
||||
Qed.
|
||||
|
||||
Lemma DLvalgetchain: forall (c:natO -m> DL) k d, c k == Val d -> exists c':natO -m> D, forall n, c(k+n) == Val (c' n).
|
||||
intros c k d chk.
|
||||
exists (makechainm chk).
|
||||
intros n.
|
||||
destruct (allvalsfromhere n chk) as [d' [chd ld]].
|
||||
refine (Oeq_trans chd _).
|
||||
apply eqDLeq.
|
||||
apply (Oeq_sym (makechainworks chk chd)).
|
||||
Qed.
|
||||
|
||||
Hint Immediate DLle_leVal.
|
||||
|
||||
(*==========================================================================
|
||||
Eta
|
||||
==========================================================================*)
|
||||
|
||||
Definition eta_m : D -m> DL.
|
||||
exists (fun (x:D) => Val x).
|
||||
unfold monotonic.
|
||||
intros;auto.
|
||||
Defined.
|
||||
|
||||
(** printing eta %\ensuremath{\eta}% *)
|
||||
Definition eta : D -c> DL.
|
||||
exists (eta_m).
|
||||
red.
|
||||
intro c.
|
||||
set (LHS := eta_m (lub c)).
|
||||
unfold DL, DL_lub.
|
||||
unfold lub.
|
||||
rewrite DL_lubn_inv.
|
||||
simpl.
|
||||
subst LHS.
|
||||
simpl.
|
||||
apply DLle_leVal.
|
||||
apply lub_le; intro n.
|
||||
setoid_replace (c n) with (makechainm (c:= eta_m @ c) (k:=0) (d := c 0) (Oeq_refl_eq (refl_equal (Val (c 0)))) n).
|
||||
apply le_lub.
|
||||
unfold makechainm.
|
||||
simpl.
|
||||
symmetry. apply makechainworks.
|
||||
simpl.
|
||||
trivial.
|
||||
Defined.
|
||||
|
||||
Add Parametric Morphism : (@eta)
|
||||
with signature (@Ole D) ++> (@Ole (DL))
|
||||
as eta_le_compat.
|
||||
intros d0 d1 Eq.
|
||||
assert (monotonic (eta)). auto. apply H.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism : (@eta)
|
||||
with signature (@Oeq D) ==> (@Oeq (DL))
|
||||
as eta_eq_compat.
|
||||
intros d0 d1 Eq. split ; auto.
|
||||
Qed.
|
||||
|
||||
Hint Resolve eta_eq_compat.
|
||||
|
||||
End Lift_cpo.
|
||||
|
||||
Implicit Arguments eta [D].
|
||||
|
||||
(** printing _BOT %\LIFTED% *)
|
||||
Notation "x '_BOT'" := (DL x) (at level 28).
|
||||
|
||||
Instance lift_pointed D : Pointed (D _BOT).
|
||||
intro D.
|
||||
exists (DL_bot D).
|
||||
intro. red. simpl.
|
||||
apply DL_bot_least.
|
||||
Defined.
|
||||
|
||||
Lemma PBot_app D E {pd:Pointed E} : forall d, @PBot _ (fun_pointed D E) d == PBot.
|
||||
intros D E pd d. split. case pd. auto.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma PBot_comp D F {pd:Pointed F} : forall E (h:E -C-> D),
|
||||
@PBot _ (fun_pointed D F) << h == @PBot _ (fun_pointed E F).
|
||||
intros D F pd E h. apply fcont_eq_intro.
|
||||
intros e. rewrite fcont_comp_simpl. rewrite PBot_app. rewrite PBot_app.
|
||||
auto.
|
||||
Qed.
|
907
papers/domains-2009/PredomProd.v
Executable file
907
papers/domains-2009/PredomProd.v
Executable file
|
@ -0,0 +1,907 @@
|
|||
(*==========================================================================
|
||||
Definition of products and associated lemmas
|
||||
==========================================================================*)
|
||||
Require Import PredomCore.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
(*==========================================================================
|
||||
Order-theoretic definitions
|
||||
==========================================================================*)
|
||||
Definition Oprod : ord -> ord -> ord.
|
||||
intros O1 O2; exists (O1 * O2)%type (fun x y => fst x <= fst y /\ snd x <= snd y); intuition.
|
||||
apply Ole_trans with a0; trivial.
|
||||
apply Ole_trans with b0; trivial.
|
||||
Defined.
|
||||
|
||||
Definition Fst (O1 O2 : ord) : Oprod O1 O2 -m> O1.
|
||||
intros O1 O2; exists (fst (A:=O1) (B:=O2)); red; simpl; intuition.
|
||||
Defined.
|
||||
|
||||
Implicit Arguments Fst [O1 O2].
|
||||
|
||||
Definition Snd (O1 O2 : ord) : Oprod O1 O2 -m> O2.
|
||||
intros O1 O2; exists (snd (A:=O1) (B:=O2)); red; simpl; intuition.
|
||||
Defined.
|
||||
|
||||
Implicit Arguments Snd [O1 O2].
|
||||
|
||||
(*
|
||||
Definition Pairr (O1 O2 : ord) : O1 -> O2 -m> Oprod O1 O2.
|
||||
intros O1 O2 x.
|
||||
exists (fun (y:O2) => (x,y)).
|
||||
red; auto.
|
||||
Defined.
|
||||
|
||||
Definition Pair (O1 O2 : ord) : O1 -m> O2 -m> Oprod O1 O2.
|
||||
intros O1 O2; exists (Pairr (O1:=O1) O2); red; auto.
|
||||
Defined.
|
||||
*)
|
||||
|
||||
Lemma Fst_simpl : forall (O1 O2 : ord) (p:Oprod O1 O2), Fst p = fst p.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma Snd_simpl : forall (O1 O2 : ord) (p:Oprod O1 O2), Snd p = snd p.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
(*
|
||||
Lemma Pair_simpl : forall (O1 O2 : ord) (x:O1)(y:O2), Pair O1 O2 x y = (x,y).
|
||||
trivial.
|
||||
Save.
|
||||
*)
|
||||
|
||||
Lemma Oprod_unique: forall O1 O2 O3 (f g:O1 -m> Oprod O2 O3),
|
||||
Fst @ f == Fst @ g -> Snd @ f == Snd @ g -> f == g.
|
||||
intros O1 O2 O3 f g ff ss.
|
||||
apply fmon_eq_intro. intros x.
|
||||
split ; auto using (fmon_eq_elim ff x), (fmon_eq_elim ss x).
|
||||
Qed.
|
||||
|
||||
(*
|
||||
Definition prod0 (D1 D2:cpo) : Oprod D1 D2 := (0: D1,0: D2).
|
||||
*)
|
||||
|
||||
Definition prod_lub (D1 D2:cpo) (f : natO -m> Oprod D1 D2) := (lub (Fst @f), lub (Snd @f)).
|
||||
|
||||
Definition Dprod : cpo -> cpo -> cpo.
|
||||
intros D1 D2; exists (Oprod D1 D2)
|
||||
(* (prod0 D1 D2) *)
|
||||
(prod_lub (D1:=D1) (D2:=D2)); unfold prod_lub; intuition.
|
||||
apply Ole_trans with (fst (fmonot c n), snd (fmonot c n)); simpl; intuition.
|
||||
apply (le_lub (Fst @ c) (n)).
|
||||
apply (le_lub (Snd @ c) (n)).
|
||||
apply Ole_trans with (fst x, snd x); simpl; intuition.
|
||||
apply lub_le; simpl; intros.
|
||||
case (H n); auto.
|
||||
apply lub_le; simpl; intros.
|
||||
case (H n); auto.
|
||||
Defined.
|
||||
|
||||
(** printing DOne %\DOne% *)
|
||||
Definition DOne : cpo := Discrete unit.
|
||||
|
||||
Lemma DOne_final: forall D (f g : D -C-> DOne), f == g.
|
||||
intros. refine (fcont_eq_intro _). intros d. case (f d) ; case (g d) ; auto.
|
||||
Qed.
|
||||
|
||||
Instance DOne_pointed : Pointed DOne.
|
||||
exists tt.
|
||||
intros d. case d. auto.
|
||||
Defined.
|
||||
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Domain-theoretic definitions
|
||||
==========================================================================*)
|
||||
|
||||
(** printing *c* %\ensuremath{\times}% *)
|
||||
Infix "*c*" := Dprod (at level 28, left associativity).
|
||||
|
||||
Lemma Dprod_eq_intro : forall (D1 D2:cpo) (p1 p2 : D1 *c* D2),
|
||||
fst p1 == fst p2 -> snd p1 == snd p2 -> p1 == p2.
|
||||
split; simpl; auto.
|
||||
Save.
|
||||
Hint Resolve Dprod_eq_intro.
|
||||
|
||||
Lemma Dprod_eq_pair : forall (D1 D2 : cpo) (x1 y1 : D1) (x2 y2 : D2),
|
||||
x1==y1 -> x2==y2 -> ((x1,x2):Dprod D1 D2) == (y1,y2).
|
||||
auto.
|
||||
Save.
|
||||
Hint Resolve Dprod_eq_pair.
|
||||
|
||||
Lemma Dprod_eq_elim_fst : forall (D1 D2 : cpo) (p1 p2 : Dprod D1 D2),
|
||||
p1==p2 -> fst p1 == fst p2.
|
||||
split; case H; simpl; intuition.
|
||||
Save.
|
||||
Hint Immediate Dprod_eq_elim_fst.
|
||||
|
||||
Lemma Dprod_eq_elim_snd : forall (D1 D2:cpo) (p1 p2: Dprod D1 D2),
|
||||
p1==p2 -> snd p1 == snd p2.
|
||||
split; case H; simpl; intuition.
|
||||
Save.
|
||||
Hint Immediate Dprod_eq_elim_snd.
|
||||
|
||||
Lemma Dprod_le_elim_snd: forall (D1 D2:cpo) (p1 p2 : Dprod D1 D2),
|
||||
p1 <= p2 -> snd p1 <= snd p2.
|
||||
simpl. intros D1 D2 p1 p2. intros x. destruct x. auto.
|
||||
Qed.
|
||||
|
||||
Hint Immediate Dprod_le_elim_snd.
|
||||
|
||||
Lemma Dprod_le_elim_fst: forall (D1 D2:cpo) (p1 p2:Dprod D1 D2),
|
||||
p1 <= p2 -> fst p1 <= fst p2.
|
||||
simpl. intros D1 D2 p1 p2 x. destruct x. auto.
|
||||
Qed.
|
||||
|
||||
Hint Immediate Dprod_le_elim_fst.
|
||||
|
||||
(** printing FST %\FST% *)
|
||||
|
||||
Definition FST (D1 D2 : cpo) : D1 *c* D2 -c> D1.
|
||||
intros; exists (@Fst D1 D2); red; intros; auto.
|
||||
Defined.
|
||||
|
||||
Implicit Arguments FST [D1 D2].
|
||||
|
||||
Add Parametric Morphism (D E : cpo) : (@FST D E)
|
||||
with signature (@Ole (Dprod D E)) ++> (@Ole D)
|
||||
as PROD_FST_le_compat.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E : cpo) : (@FST D E)
|
||||
with signature (@Oeq (Dprod D E)) ==> (@Oeq D)
|
||||
as PROD_FST_eq_compat.
|
||||
intros ; split ; auto.
|
||||
Qed.
|
||||
|
||||
(** printing SND %\SND% *)
|
||||
Definition SND (D1 D2 : cpo) : D1 *c* D2 -c> D2.
|
||||
intros; exists (@Snd D1 D2); red; intros; auto.
|
||||
Defined.
|
||||
|
||||
Implicit Arguments SND [D1 D2].
|
||||
|
||||
Add Parametric Morphism (D E : cpo) : (@SND D E)
|
||||
with signature (@Ole (Dprod D E)) ++> (@Ole E)
|
||||
as PROD_SND_le_compat.
|
||||
intros ; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E : cpo) : (@SND D E)
|
||||
with signature (@Oeq (Dprod D E)) ==> (@Oeq E)
|
||||
as PROD_SND_eq_compat.
|
||||
intros ; split ; auto.
|
||||
Qed.
|
||||
|
||||
Definition Pairr (O1 O2 : ord) : O1 -> O2 -m> Oprod O1 O2.
|
||||
intros O1 O2 x.
|
||||
exists (fun (y:O2) => (x,y)).
|
||||
red; auto.
|
||||
Defined.
|
||||
|
||||
Definition Pair (O1 O2 : ord) : O1 -m> O2 -m> Oprod O1 O2.
|
||||
intros O1 O2; exists (Pairr (O1:=O1) O2); red; auto.
|
||||
Defined.
|
||||
|
||||
Lemma Pair_continuous2 : forall (D1 D2 : cpo), continuous2 (D3:=Dprod D1 D2) (Pair D1 D2).
|
||||
red; intros; auto.
|
||||
Save.
|
||||
|
||||
Definition PAIR (D1 D2:cpo) : D1 -c> D2 -C-> Dprod D1 D2
|
||||
:= continuous2_cont (Pair_continuous2 (D1:=D1) (D2:=D2)).
|
||||
|
||||
Implicit Arguments PAIR [D1 D2].
|
||||
|
||||
Add Parametric Morphism (D0 D1:cpo) : (@pair D0 D1)
|
||||
with signature (@Ole D0) ++> (@Ole D1) ++> (@Ole (Dprod D0 D1))
|
||||
as pair_le_compat.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D0 D1:cpo) : (@pair D0 D1)
|
||||
with signature (@Oeq D0) ==> (@Oeq D1) ==> (@Oeq (Dprod D0 D1))
|
||||
as pair_eq_compat.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma Dprod_unique: forall D E F (f g : F -C-> D *c* E),
|
||||
FST << f == FST << g ->
|
||||
SND << f == SND << g ->
|
||||
f == g.
|
||||
Proof.
|
||||
intros D E F f g ff ss.
|
||||
refine (Oprod_unique ff ss).
|
||||
Qed.
|
||||
|
||||
Lemma FST_simpl : forall (D1 D2 :cpo) (p:Dprod D1 D2), FST p = Fst p.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma SND_simpl : forall (D1 D2 :cpo) (p:Dprod D1 D2), SND p = Snd p.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Definition Prod_fun (O1 O2 O3:ord)(f:O1-m>O2)(g:O1-m>O3) : O1 -m> Oprod O2 O3.
|
||||
intros O1 O2 O3 f g. exists (fun p => (f p, g p)). auto.
|
||||
Defined.
|
||||
|
||||
Lemma Prod_fun_simpl (O1 O2 O3:ord)(f:O1-m>O2)(g:O1-m>O3) x : Prod_fun f g x = (f x,g x).
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
|
||||
Definition PROD_fun (D1 D2 D3:cpo)(f:D1-c>D2)(g:D1-c>D3) : D1 -c> D2 *c* D3.
|
||||
intros D1 D2 D3 f g. exists (Prod_fun (fcontit f) (fcontit g)).
|
||||
unfold continuous.
|
||||
intros c.
|
||||
simpl. split.
|
||||
apply Ole_trans with (y:= f (lub c)) ; auto.
|
||||
rewrite (fcont_continuous f) ; auto.
|
||||
apply Ole_trans with (y:= g (lub c)) ; auto.
|
||||
rewrite (fcont_continuous g) ; auto.
|
||||
Defined.
|
||||
|
||||
|
||||
Lemma FST_PAIR_simpl : forall (D1 D2 :cpo) (p1:D1) (p2:D2),
|
||||
FST (PAIR p1 p2) = p1.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma SND_PAIR_simpl : forall (D1 D2 :cpo) (p1:D1) (p2:D2),
|
||||
SND (PAIR p1 p2) = p2.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma PAIR_simpl : forall (D1 D2 :cpo) (p1:D1) (p2:D2), PAIR p1 p2 = Pair D1 D2 p1 p2.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
(** printing <| %\ensuremath{\langle}% *)
|
||||
(** printing |> %\ensuremath{\rangle}% *)
|
||||
Notation "'<|' f , g '|>'" := (PROD_fun f g) (at level 30).
|
||||
|
||||
Lemma PROD_fun_simpl (D E F : cpo) (f : D -C-> E) (g : D -C-> F) d :
|
||||
<| f, g |> d = (f d , g d).
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma FST_PROD_fun: forall D0 D1 D2 (f:D0 -C-> D1) (g:D0 -C-> D2), FST << PROD_fun f g == f.
|
||||
intros. apply fcont_eq_intro ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma SND_PROD_fun: forall D0 D1 D2 (f:D0 -C-> D1) (g:D0 -C-> D2), SND << PROD_fun f g == g.
|
||||
intros. apply fcont_eq_intro ; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D0 D1 D2:cpo) : (@PROD_fun D0 D1 D2)
|
||||
with signature (@Ole (D0 -c> D1)) ++> (@Ole (D0 -c> D2)) ++> (@Ole (D0 -c> @Dprod D1 D2))
|
||||
as PROD_fun_le_compat.
|
||||
intros f g feq s r seq.
|
||||
simpl.
|
||||
intros x.
|
||||
assert (fd:=fcont_le_elim feq x). assert (fe := fcont_le_elim seq x).
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D0 D1 D2:cpo) : (@PROD_fun D0 D1 D2)
|
||||
with signature (@Oeq (D0 -c> D1)) ==> (@Oeq (D0 -c> D2)) ==> (@Oeq (D0 -c> @Dprod D1 D2))
|
||||
as PROD_fun_eq_compat.
|
||||
intros f g feq s r geq.
|
||||
destruct feq as [fle1 fle2].
|
||||
destruct geq as [gle1 gle2].
|
||||
split ; auto.
|
||||
Qed.
|
||||
|
||||
Definition PROD_map (D1 D2 D3 D4:cpo)(f:D1-c>D3)(g:D2-c>D4) :
|
||||
D1 *c* D2 -c> D3 *c* D4 := PROD_fun (f << FST) (g << SND).
|
||||
|
||||
Implicit Arguments PROD_map [D1 D2 D3 D4].
|
||||
|
||||
(** printing *f* %\crossf% *)
|
||||
Infix "*f*" := PROD_map (at level 28).
|
||||
|
||||
Lemma PROD_fun_compl (D E F G : cpo) (f : D -C-> E) (g : D -C-> F) (h : G -C-> D) :
|
||||
<| f, g|> << h == <| f << h , g << h |>.
|
||||
intros D E F G f g h.
|
||||
simpl. apply (@Dprod_unique E F G).
|
||||
rewrite FST_PROD_fun. rewrite <-fcont_comp_assoc. rewrite FST_PROD_fun. auto.
|
||||
rewrite SND_PROD_fun. rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. auto.
|
||||
Qed.
|
||||
|
||||
Lemma PROD_fun_compr (D E F G H : cpo) (f : H -c> D) (g : D -c> E) (h : H -c> F) (k : F -c> G) :
|
||||
<| g << f, k << h |> == g *f* k << <| f, h |>.
|
||||
intros D E F G H f g h k. apply (@Dprod_unique E G H).
|
||||
rewrite FST_PROD_fun. rewrite <-fcont_comp_assoc. unfold PROD_map.
|
||||
rewrite FST_PROD_fun. rewrite fcont_comp_assoc. rewrite FST_PROD_fun. auto.
|
||||
rewrite SND_PROD_fun. rewrite <-fcont_comp_assoc. unfold PROD_map.
|
||||
rewrite SND_PROD_fun. rewrite fcont_comp_assoc. rewrite SND_PROD_fun. auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E F G:cpo) : (@PROD_map D E F G)
|
||||
with signature (@Ole (D -c> F)) ++> (@Ole (E -c> G)) ++> (@Ole (@Dprod D E -c> @Dprod F G))
|
||||
as PROD_map_le_compat.
|
||||
intros f g feq s r seq.
|
||||
simpl.
|
||||
intros x. case x. clear x. intros d e.
|
||||
assert (fd:=fcont_le_elim feq d). assert (fe := fcont_le_elim seq e).
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E F G:cpo) : (@PROD_map D E F G)
|
||||
with signature (@Oeq (D -c> F)) ==> (@Oeq (E -c> G)) ==> (@Oeq (@Dprod D E -c> @Dprod F G))
|
||||
as PROD_map_eq_compat.
|
||||
intros f g feq s r geq.
|
||||
destruct feq as [fle1 fle2].
|
||||
destruct geq as [gle1 gle2].
|
||||
split ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma PROD_map_simpl : forall (D1 D2 D3 D4 : cpo) (f : D1 -c> D3) (g : D2 -c> D4) (p : D1 *c* D2),
|
||||
(f *f* g) p = pair (f (fst p)) (g (snd p)).
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Definition PROD_Map1 (D E F G:cpo) : (D -C-> F) -> (E -C-> G) -M-> Dprod D E -C-> Dprod F G.
|
||||
intros. exists (@PROD_map D E F G X). unfold monotonic.
|
||||
intros f0 f1 leq. intros x. simpl ; auto.
|
||||
Defined.
|
||||
|
||||
Definition PROD_Map (D E F G:cpo) : (D -C-> F) -M-> (E -C-> G) -M-> Dprod D E -C-> Dprod F G.
|
||||
intros. exists (@PROD_Map1 D E F G). unfold monotonic.
|
||||
intros f0 f1 leq. intros x. simpl ; auto.
|
||||
Defined.
|
||||
|
||||
Lemma PROD_Map_continuous2 : forall D E F G, continuous2 (PROD_Map D E F G).
|
||||
intros. unfold continuous2. intros c1 c2.
|
||||
intros p. auto.
|
||||
Save.
|
||||
|
||||
Definition PROD_MAP (D E F G:cpo) : (D -C-> F) -C-> (E -C-> G) -C-> Dprod D E -C-> Dprod F G :=
|
||||
(continuous2_cont (@PROD_Map_continuous2 D E F G)).
|
||||
|
||||
Definition SWAP D E := <| @SND D E, FST |>.
|
||||
Implicit Arguments SWAP [D E].
|
||||
|
||||
Instance prod_pointed A B { pa : Pointed A} {pb : Pointed B} : Pointed (A *c* B).
|
||||
intros.
|
||||
destruct pa as [pa Pa0].
|
||||
destruct pb as [pb Pb0].
|
||||
exists (pa,pb).
|
||||
intro.
|
||||
auto.
|
||||
Defined.
|
||||
|
||||
Lemma PROD_map_compl D0 D1 D2 D3 D4 (f:D0 -c> D1) (g:D1 -c> D2) (h:D3 -c> D4) : (g << f) *f* h == (g *f* ID) << (f *f* h).
|
||||
intros ; refine (Dprod_unique _ _). unfold PROD_map. rewrite FST_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite FST_PROD_fun.
|
||||
auto. unfold PROD_map. rewrite SND_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite SND_PROD_fun.
|
||||
rewrite fcont_comp_unitL.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma PROD_map_compr D0 D1 D2 D3 D4 (f:D0 -c> D1) (g:D1 -c> D2) (h:D3 -c> D4) : h *f* (g << f) == (ID *f* g) << (h *f* f).
|
||||
intros ; refine (Dprod_unique _ _). unfold PROD_map. rewrite FST_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite FST_PROD_fun.
|
||||
rewrite fcont_comp_unitL. auto.
|
||||
unfold PROD_map. rewrite SND_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite SND_PROD_fun.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Definition curry (D1 D2 D3 : cpo) (f:Dprod D1 D2 -c> D3) : D1 -c> (D2-C->D3) :=
|
||||
fcont_COMP D1 (D2-C->Dprod D1 D2) (D2-C->D3)
|
||||
(fcont_COMP D2 (Dprod D1 D2) D3 f) (PAIR).
|
||||
|
||||
Definition EV D0 D1 := ((AP D0 D1) @2_ FST) SND.
|
||||
|
||||
Implicit Arguments EV [D0 D1].
|
||||
|
||||
Lemma EV_simpl D0 D1 (f : D0 -c> D1) a : EV (f,a) = f a.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism D0 D1 : (@EV D0 D1)
|
||||
with signature (@Ole (Dprod (D0 -C-> D1) D0)) ++> (@Ole D1)
|
||||
as EV_le_compat.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism D E : (@EV D E)
|
||||
with signature (@Oeq (Dprod (D -C-> E) D)) ==> (@Oeq E)
|
||||
as EV_eq_compat.
|
||||
split ; auto.
|
||||
Qed.
|
||||
|
||||
Definition CURRY (D0 D1 D2:cpo) : (Dprod D0 D1 -C-> D2) -c> (D0 -C-> D1 -C-> D2) :=
|
||||
curry (curry (EV << <| FST << FST, <| SND << FST, SND |> |>)).
|
||||
|
||||
Implicit Arguments CURRY [D0 D1 D2].
|
||||
|
||||
Lemma CURRY_simpl : forall (D1 D2 D3 : cpo) (f:Dprod D1 D2 -C-> D3),
|
||||
CURRY f == curry f.
|
||||
intros. unfold CURRY.
|
||||
simpl. apply fcont_eq_intro. intros x. simpl. apply fcont_eq_intro. intros y. auto.
|
||||
Save.
|
||||
|
||||
Lemma curry_com: forall D E F (f : D *c* E -C-> F),
|
||||
f == EV << (PROD_map (curry f) ID).
|
||||
intros. simpl. apply fcont_eq_intro. intros d. case d. clear d.
|
||||
intros d e. auto.
|
||||
Qed.
|
||||
|
||||
Lemma curry_unique: forall D E F (f:D -c> E -C-> F) g,
|
||||
g == EV << (PROD_map f ID) -> f == curry g.
|
||||
intros. apply fcont_eq_intro. intros a. simpl. apply fcont_eq_intro. intros b.
|
||||
assert (xx:=fcont_eq_elim H (a,b)). auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D0 D1 D2 : cpo) : (@curry D0 D1 D2)
|
||||
with signature (@Ole (Dprod D0 D1 -c> D2)) ++> (@Ole (D0 -c> (D1 -C-> D2)))
|
||||
as curry_le_compat.
|
||||
intros f g lf.
|
||||
simpl. intros x0 x1. auto.
|
||||
Qed.
|
||||
|
||||
Hint Resolve curry_le_compat_Morphism.
|
||||
|
||||
Add Parametric Morphism (D0 D1 D2 : cpo) : (@curry D0 D1 D2)
|
||||
with signature (@Oeq (Dprod D0 D1 -c> D2)) ==> (@Oeq (D0 -c> (D1 -C-> D2)))
|
||||
as curry_eq_compat.
|
||||
intros f1 f2 Eqf.
|
||||
destruct Eqf as [fle1 fle2].
|
||||
split.
|
||||
apply (curry_le_compat fle1).
|
||||
apply (curry_le_compat fle2).
|
||||
Qed.
|
||||
|
||||
Hint Resolve curry_eq_compat_Morphism.
|
||||
|
||||
Lemma curry_simpl: forall D E F (f:(Dprod E F) -C-> D) x y, curry f x y = f (x,y).
|
||||
intros D E F f x y.
|
||||
unfold curry.
|
||||
repeat (try (rewrite fcont_COMP_simpl) ; try (rewrite fcont_comp_simpl) ; try (rewrite PAIR_simpl)).
|
||||
simpl. auto.
|
||||
Qed.
|
||||
|
||||
Lemma Dext (D0 D1 D2:cpo) (f g : D0 -c> D1 -C-> D2) : EV << (f *f* ID) == EV << (g *f* ID) -> f == g.
|
||||
intros D0 D1 D2 f g C.
|
||||
rewrite (curry_unique (Oeq_refl (EV << (f *f* ID)))).
|
||||
rewrite (curry_unique (Oeq_refl (EV << (g *f* ID)))).
|
||||
apply (curry_eq_compat C).
|
||||
Qed.
|
||||
|
||||
Definition uncurry D0 D1 D2 :=
|
||||
curry (EV << <| EV << <| @FST (D0 -C-> D1 -C-> D2) _, FST << SND |>,
|
||||
SND << SND |>).
|
||||
|
||||
Implicit Arguments uncurry [D0 D1 D2].
|
||||
|
||||
Lemma uncurry_simpl: forall D E F (f:D -c> E -C-> F) p1 p2, uncurry f (p1,p2) = f p1 p2.
|
||||
intros ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma PROD_fun_comp (D0 D1 D2 D3:cpo) (f:D0 -c> D1) (g:D0 -c> D2) (h:D3 -c> D0) : <| f,g |> << h == <| f << h, g << h |>.
|
||||
intros.
|
||||
refine (Dprod_unique _ _); rewrite <- fcont_comp_assoc.
|
||||
rewrite FST_PROD_fun ; rewrite FST_PROD_fun ; auto.
|
||||
rewrite SND_PROD_fun ; rewrite SND_PROD_fun ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma curry_uncurry (D0 D1 D2:cpo) : @CURRY D0 D1 D2 << uncurry == ID.
|
||||
intros D0 D1 D2.
|
||||
|
||||
assert (forall D0 D1 D2 D3 D4 (f:D0 -c> D1) (g:D1 -C-> D2) (h:D3 -c> D4), (g << f) *f* h == (g *f* ID) << (f *f* h)).
|
||||
intros. refine (Dprod_unique _ _) ; unfold PROD_map. rewrite FST_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite FST_PROD_fun. auto.
|
||||
rewrite SND_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite SND_PROD_fun.
|
||||
rewrite fcont_comp_unitL. auto.
|
||||
|
||||
apply Dext. rewrite H.
|
||||
rewrite <- fcont_comp_assoc.
|
||||
unfold CURRY. rewrite <- curry_com.
|
||||
apply Dext.
|
||||
rewrite H.
|
||||
rewrite <- fcont_comp_assoc.
|
||||
rewrite <- curry_com.
|
||||
unfold uncurry.
|
||||
rewrite fcont_comp_assoc.
|
||||
assert ((<| FST << FST, <| SND << FST, SND |> |> <<
|
||||
(curry (D1:=D0 -C-> D1 -C-> D2) (D2:=D0 *c* D1) (D3:=D2)
|
||||
(EV << <| EV << <| FST, FST << SND |>, SND << SND |>) *f*
|
||||
@ID D0) *f* @ID D1) ==
|
||||
(curry (EV << <| EV << <| FST, FST << SND |>, SND << SND |>) *f* (ID *f* ID)) <<
|
||||
<| FST << FST, <| SND << FST, SND |> |>).
|
||||
refine (Dprod_unique _ _) ; unfold PROD_map.
|
||||
rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. rewrite fcont_comp_assoc.
|
||||
rewrite FST_PROD_fun. rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite FST_PROD_fun. auto.
|
||||
rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun.
|
||||
refine (Dprod_unique _ _). rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun.
|
||||
rewrite fcont_comp_assoc. rewrite FST_PROD_fun. rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun.
|
||||
repeat (rewrite <- fcont_comp_assoc). rewrite FST_PROD_fun.
|
||||
repeat (rewrite fcont_comp_unitL).
|
||||
repeat (rewrite fcont_comp_assoc). rewrite SND_PROD_fun. rewrite FST_PROD_fun. auto.
|
||||
rewrite <- fcont_comp_assoc. repeat (rewrite SND_PROD_fun).
|
||||
repeat (rewrite <- fcont_comp_assoc). rewrite SND_PROD_fun. rewrite fcont_comp_assoc.
|
||||
rewrite SND_PROD_fun. rewrite fcont_comp_assoc.
|
||||
rewrite SND_PROD_fun. auto.
|
||||
rewrite H0. clear H0. rewrite <- fcont_comp_assoc.
|
||||
assert (@ID D0 *f* @ID D1 == ID) as Ide. refine (Dprod_unique _ _) ; unfold PROD_map.
|
||||
rewrite FST_PROD_fun. rewrite fcont_comp_unitR. rewrite fcont_comp_unitL. auto.
|
||||
rewrite SND_PROD_fun. rewrite fcont_comp_unitR. rewrite fcont_comp_unitL. auto.
|
||||
rewrite (PROD_map_eq_compat (Oeq_refl _) Ide). rewrite <- curry_com.
|
||||
repeat (rewrite fcont_comp_assoc).
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
refine (Dprod_unique _ _) ; unfold PROD_map.
|
||||
rewrite <- fcont_comp_assoc. repeat (rewrite FST_PROD_fun).
|
||||
repeat (rewrite fcont_comp_assoc).
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
refine (Dprod_unique _ _). repeat (rewrite <- fcont_comp_assoc). repeat (rewrite FST_PROD_fun).
|
||||
rewrite fcont_comp_unitL. auto. repeat (rewrite <- fcont_comp_assoc). repeat (rewrite SND_PROD_fun).
|
||||
rewrite fcont_comp_unitL. rewrite fcont_comp_assoc. rewrite SND_PROD_fun. rewrite FST_PROD_fun. auto.
|
||||
repeat (rewrite <- fcont_comp_assoc). repeat (rewrite SND_PROD_fun).
|
||||
rewrite fcont_comp_assoc. repeat (rewrite SND_PROD_fun). rewrite fcont_comp_unitL. auto.
|
||||
Qed.
|
||||
|
||||
(*
|
||||
Lemma uncurry_curry (D0 D1 D2:cpo) : uncurry << @CURRY D0 D1 D2 == ID.
|
||||
intros.
|
||||
apply Dext. unfold uncurry.
|
||||
assert (forall D0 D1 D2 D3 D4 (f:D0 -c> D1) (g:D1 -C-> D2) (h:D3 -c> D4), (g << f) *f* h == (g *f* ID) << (f *f* h)).
|
||||
intros. refine (Dprod_unique _ _) ; unfold PROD_map. rewrite FST_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite FST_PROD_fun. auto.
|
||||
rewrite SND_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. repeat (rewrite fcont_comp_assoc). rewrite SND_PROD_fun.
|
||||
rewrite fcont_comp_unitL. auto.
|
||||
|
||||
rewrite H. rewrite <- fcont_comp_assoc. rewrite <- curry_com. unfold CURRY.
|
||||
|
||||
*)
|
||||
|
||||
Add Parametric Morphism (D0 D1 D2:cpo) : (@uncurry D0 D1 D2)
|
||||
with signature (@Oeq (D0 -c> (D1 -C-> D2))) ==> (@Oeq ((Dprod D0 D1) -c> D2))
|
||||
as uncurry_eq_compat.
|
||||
intros. apply fcont_eq_intro ; auto.
|
||||
Qed.
|
||||
Hint Resolve uncurry_eq_compat_Morphism.
|
||||
|
||||
Add Parametric Morphism (D0 D1 D2:cpo) : (@uncurry D0 D1 D2)
|
||||
with signature (@Ole (D0 -c> (D1 -C-> D2))) ++> (@Ole ((Dprod D0 D1) -c> D2))
|
||||
as uncurry_le_compat.
|
||||
intros. apply fcont_le_intro ; auto.
|
||||
Qed.
|
||||
Hint Resolve uncurry_le_compat_Morphism.
|
||||
|
||||
Lemma Uncurry_simpl : forall (D1 D2 D3 : cpo) (f:D1 -c> (D2-C->D3)) (p:Dprod D1 D2),
|
||||
uncurry f p = f (fst p) (snd p).
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma FST_SWAP : forall D E F (f:D -c> F), f << FST << SWAP == f << SND (D1:=E).
|
||||
intros; apply fcont_eq_intro; auto.
|
||||
Qed.
|
||||
|
||||
Lemma SND_SWAP : forall D E F (f:E -c> F), f << SND << SWAP == f << FST (D2:=D).
|
||||
intros; apply fcont_eq_intro. auto.
|
||||
Qed.
|
||||
|
||||
Lemma SND_PAIR : forall D E F (f:E -c> F) x, f << SND << PAIR (D1:=D) x == f.
|
||||
intros; apply fcont_eq_intro. auto.
|
||||
Qed.
|
||||
|
||||
Lemma FST_PAIR : forall D E F (f:D -c> F) (x:D), f << FST << PAIR x == @K E _ (f x).
|
||||
intros; apply fcont_eq_intro. auto.
|
||||
Qed.
|
||||
|
||||
Lemma FST_PROD_fun_aux : forall C D E F (f:C -c> D) (g:C -c> E) (k:D -c> F), k << FST << PROD_fun f g == k << f.
|
||||
intros; apply fcont_eq_intro. auto.
|
||||
Qed.
|
||||
|
||||
Lemma SND_PROD_fun_aux : forall C D E F (f:C -c> D) (g:C -c> E) (k:E -c> F), k << SND << PROD_fun f g == k << g.
|
||||
intros; apply fcont_eq_intro. auto.
|
||||
Qed.
|
||||
|
||||
Lemma PROD_map_ID_ID :
|
||||
forall C D, PROD_map (D1 := C) (D2 := D) ID ID == ID.
|
||||
intros C D.
|
||||
apply fcont_eq_intro. auto.
|
||||
Qed.
|
||||
|
||||
Lemma PROD_map_PROD_map : forall A B C D E F (f:A -c> B) (g : C -c> D) (h :B -c> E) (i :D -c> F),
|
||||
PROD_map h i << PROD_map f g == PROD_map (h << f) (i << g).
|
||||
intros. apply fcont_eq_intro. auto.
|
||||
Qed.
|
||||
|
||||
Lemma Curry_EV : forall C D E (f:C -c> D -C-> E) , curry (EV << PROD_map f ID) == f.
|
||||
intros C D E f.
|
||||
symmetry.
|
||||
apply curry_unique. trivial.
|
||||
Qed.
|
||||
|
||||
Lemma EV_Curry : forall C D E (f:Dprod C D -c> E) , EV << PROD_map (curry f) ID == f.
|
||||
intros C D E f.
|
||||
rewrite <- curry_com. trivial.
|
||||
Qed.
|
||||
|
||||
Lemma Curry_comp : forall A B C D (f:Dprod A B -c> C) (g:D -c> A), curry f << g == curry (f << PROD_map g ID).
|
||||
intros A B C D f g.
|
||||
assert (H := curry_com f). rewrite H.
|
||||
rewrite fcont_comp_assoc.
|
||||
rewrite PROD_map_PROD_map.
|
||||
rewrite Curry_EV. rewrite fcont_comp_unitL.
|
||||
rewrite Curry_EV. trivial.
|
||||
Qed.
|
||||
|
||||
|
||||
|
||||
|
||||
|
||||
(** ** Indexed product of cpo's *)
|
||||
|
||||
Definition Oprodi (I:Type)(O:I->ord) : ord.
|
||||
intros; exists (forall i:I, O i) (fun p1 p2 => forall i:I, p1 i <= p2 i); intros; auto.
|
||||
apply Ole_trans with (y i); trivial.
|
||||
Defined.
|
||||
|
||||
Lemma Oprodi_eq_intro : forall (I:Type)(O:I->ord) (p q : Oprodi O), (forall i, p i == q i) -> p==q.
|
||||
intros; apply Ole_antisym; intro i; auto.
|
||||
Save.
|
||||
|
||||
Lemma Oprodi_eq_elim : forall (I:Type)(O:I->ord) (p q : Oprodi O), p==q -> forall i, p i == q i.
|
||||
intros; apply Ole_antisym; case H; auto.
|
||||
Save.
|
||||
|
||||
Definition Proj (I:Type)(O:I->ord) (i:I) : Oprodi O -m> O i.
|
||||
intros.
|
||||
exists (fun (x:Oprodi O) => x i); red; intuition.
|
||||
Defined.
|
||||
|
||||
Lemma Proj_simpl : forall (I:Type)(O:I->ord) (i:I) (x:Oprodi O),
|
||||
Proj O i x = x i.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Definition Dprodi (I:Type)(D:I->cpo) : cpo.
|
||||
intros; exists (Oprodi D) (fun (f : natO -m> Oprodi D) (i:I) => lub (Proj D i @ f));
|
||||
intros; simpl; intros; auto.
|
||||
apply (le_lub (Proj (fun x : I => D x) i @ c) (n)).
|
||||
apply lub_le; simpl; intros.
|
||||
apply (H n i).
|
||||
Defined.
|
||||
|
||||
Lemma Dprodi_lub_simpl : forall (I:Type)(Di:I->cpo)(h:natO-m>Dprodi Di)(i:I),
|
||||
lub h i = lub (c:=Di i) (Proj Di i @ h).
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma Dprodi_eq_intro : forall I O (p1 p2:@Dprodi I O),
|
||||
(forall i, Proj O i p1 == Proj O i p2) -> p1 == p2.
|
||||
intros I O p1 p2 C.
|
||||
split ; intro i ; specialize (C i) ; simpl in C; auto.
|
||||
Qed.
|
||||
|
||||
Lemma Dprodi_eq_elim : forall I O i (p1 p2: @Dprodi I O),
|
||||
p1 == p2 -> Proj O i p1 == Proj O i p2.
|
||||
intros. simpl. simpl in H. destruct H as [X Y].
|
||||
split ; auto.
|
||||
Qed.
|
||||
|
||||
Hint Resolve Dprodi_eq_intro.
|
||||
Hint Immediate Dprodi_eq_elim.
|
||||
|
||||
Lemma Dprodi_le_intro : forall I O (p1 p2:@Dprodi I O),
|
||||
(forall i, Proj O i p1 <= Proj O i p2) -> p1 <= p2.
|
||||
intros I O p1 p2 C.
|
||||
intro i ; specialize (C i) ; simpl in C; auto.
|
||||
Qed.
|
||||
|
||||
Lemma Dprodi_le_elim : forall I O i (p1 p2: @Dprodi I O),
|
||||
p1 <= p2 -> Proj O i p1 <= Proj O i p2.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Hint Resolve Dprodi_le_intro.
|
||||
Hint Immediate Dprodi_le_elim.
|
||||
|
||||
Lemma Proj_cont : forall (I:Type)(Di:I->cpo) (i:I),
|
||||
continuous (D1:=Dprodi Di) (D2:=Di i) (Proj Di i).
|
||||
red; intros; simpl; trivial.
|
||||
Save.
|
||||
|
||||
Definition PROJ (I:Type)(Di:I->cpo) (i:I) : Dprodi Di -c> Di i :=
|
||||
mk_fconti (Proj_cont (Di:=Di) i).
|
||||
|
||||
Lemma PROJ_simpl : forall (I:Type)(Di:I->cpo) (i:I)(d:Dprodi Di),
|
||||
PROJ Di i d = d i.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma Dprodi_unique: forall I O D (f g : D -C-> @Dprodi I O),
|
||||
(forall i, PROJ O i << f == PROJ O i << g) -> f == g.
|
||||
intros. simpl.
|
||||
apply fcont_eq_intro. intros x. simpl.
|
||||
refine (Dprodi_eq_intro _).
|
||||
intros i. specialize (H i). assert (h:=fcont_eq_elim H x).
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Definition Prodi_fun I O D (f:forall i:I, D -m> O i) : D -m> Oprodi O.
|
||||
intros. exists (fun d => fun i => f i d).
|
||||
unfold monotonic. intros d0 d1 deq. simpl. intros i.
|
||||
assert (monotonic (f i)) as M by auto.
|
||||
apply M. auto.
|
||||
Defined.
|
||||
|
||||
Definition PRODI_fun I O D (f:forall i:I, D -c> O i) : D -c> Dprodi O.
|
||||
intros. exists (@Prodi_fun I _ D (fun i => fcontit (f i))).
|
||||
unfold continuous. intros c. simpl.
|
||||
intros i. assert (continuous (fcontit (f i))) as C by auto.
|
||||
unfold continuous in C. rewrite C.
|
||||
refine (lub_le_compat _). auto.
|
||||
Defined.
|
||||
|
||||
Lemma PROJ_com_point: forall D I O (f:forall i, D -C-> O i) (i:I) d,
|
||||
PROJ O i (PRODI_fun f d) = f i d.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma PROJ_com: forall D I O (f:forall i, D -C-> O i) (i:I),
|
||||
PROJ O i << PRODI_fun f == f i.
|
||||
auto using fcont_eq_intro.
|
||||
Qed.
|
||||
|
||||
Definition PRODI_map I O1 O2 (f:forall i:I, O1 i -c> O2 i) :=
|
||||
PRODI_fun (fun i => f i << PROJ O1 i).
|
||||
|
||||
Lemma Dprodi_continuous : forall (D:cpo)(I:Type)(Di:I->cpo)
|
||||
(f:D -m> Dprodi Di), (forall i, continuous (Proj Di i @ f)) ->
|
||||
continuous f.
|
||||
red; intros; intro i.
|
||||
apply Ole_trans with (lub (c:=Di i) ((Proj Di i @ f) @ c)); auto.
|
||||
exact (H i c).
|
||||
Save.
|
||||
|
||||
Definition Dprodi_lift : forall (I J:Type)(Di:I->cpo)(f:J->I),
|
||||
Dprodi Di -m> Dprodi (fun j => Di (f j)).
|
||||
intros.
|
||||
exists (fun (p:Dprodi Di) j => p (f j)); red; auto.
|
||||
Defined.
|
||||
|
||||
Lemma Dprodi_lift_simpl : forall (I J:Type)(Di:I->cpo)(f:J->I)(p:Dprodi Di),
|
||||
Dprodi_lift Di f p = fun j => p (f j).
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma Dprodi_lift_cont : forall (I J:Type)(Di:I->cpo)(f:J->I),
|
||||
continuous (Dprodi_lift Di f).
|
||||
intros; apply Dprodi_continuous; red; simpl; intros; auto.
|
||||
Save.
|
||||
|
||||
Definition DLIFTi (I J:Type)(Di:I->cpo)(f:J->I) : Dprodi Di -c> Dprodi (fun j => Di (f j))
|
||||
:= mk_fconti (Dprodi_lift_cont (Di:=Di) f).
|
||||
|
||||
Definition Dmapi : forall (I:Type)(Di Dj:I->cpo)(f:forall i, Di i -m> Dj i),
|
||||
Dprodi Di -m> Dprodi Dj.
|
||||
intros; exists (fun p i => f i (p i)); red; auto.
|
||||
Defined.
|
||||
|
||||
Lemma Dmapi_simpl : forall (I:Type)(Di Dj:I->cpo)(f:forall i, Di i -m> Dj i) (p:Dprodi Di) (i:I),
|
||||
Dmapi f p i = f i (p i).
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
|
||||
(* Stuff
|
||||
(** *** Particular cases with one or two elements *)
|
||||
|
||||
Section Product2.
|
||||
|
||||
Definition I2 := bool.
|
||||
Variable DI2 : bool -> cpo.
|
||||
|
||||
Definition DP1 := DI2 true.
|
||||
Definition DP2 := DI2 false.
|
||||
|
||||
Definition PI1 : Dprodi DI2 -c> DP1 := PROJ DI2 true.
|
||||
Definition pi1 (d:Dprodi DI2) := PI1 d.
|
||||
|
||||
Definition PI2 : Dprodi DI2 -c> DP2 := PROJ DI2 false.
|
||||
Definition pi2 (d:Dprodi DI2) := PI2 d.
|
||||
|
||||
Definition pair2 (d1:DP1) (d2:DP2) : Dprodi DI2 := bool_rect DI2 d1 d2.
|
||||
|
||||
Lemma pair2_le_compat : forall (d1 d'1:DP1) (d2 d'2:DP2), d1 <= d'1 -> d2 <= d'2
|
||||
-> pair2 d1 d2 <= pair2 d'1 d'2.
|
||||
intros; intro b; case b; simpl; auto.
|
||||
Save.
|
||||
|
||||
Definition Pair2 : DP1 -m> DP2 -m> Dprodi DI2 := le_compat2_mon pair2_le_compat.
|
||||
|
||||
Definition PAIR2 : DP1 -c> DP2 -C-> Dprodi DI2.
|
||||
apply continuous2_cont with (f:=Pair2).
|
||||
red; intros; intro b.
|
||||
case b; simpl; apply lub_le_compat; auto.
|
||||
Defined.
|
||||
|
||||
Lemma PAIR2_simpl : forall (d1:DP1) (d2:DP2), PAIR2 d1 d2 = Pair2 d1 d2.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma Pair2_simpl : forall (d1:DP1) (d2:DP2), Pair2 d1 d2 = pair2 d1 d2.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma pi1_simpl : forall (d1: DP1) (d2:DP2), pi1 (pair2 d1 d2) = d1.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma pi2_simpl : forall (d1: DP1) (d2:DP2), pi2 (pair2 d1 d2) = d2.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Definition DI2_map (f1 : DP1 -c> DP1) (f2:DP2 -c> DP2)
|
||||
: Dprodi DI2 -c> Dprodi DI2 :=
|
||||
DMAPi (bool_rect (fun b:bool => DI2 b -c>DI2 b) f1 f2).
|
||||
|
||||
Lemma Dl2_map_eq : forall (f1 : DP1 -c> DP1) (f2:DP2 -c> DP2) (d:Dprodi DI2),
|
||||
DI2_map f1 f2 d == pair2 (f1 (pi1 d)) (f2 (pi2 d)).
|
||||
intros; simpl; apply Oprodi_eq_intro; intro b; case b; trivial.
|
||||
Save.
|
||||
End Product2.
|
||||
Hint Resolve Dl2_map_eq.
|
||||
|
||||
Section Product1.
|
||||
Definition I1 := unit.
|
||||
Variable D : cpo.
|
||||
|
||||
Definition DI1 (_:unit) := D.
|
||||
Definition PI : Dprodi DI1 -c> D := PROJ DI1 tt.
|
||||
Definition pi (d:Dprodi DI1) := PI d.
|
||||
|
||||
Definition pair1 (d:D) : Dprodi DI1 := unit_rect DI1 d.
|
||||
|
||||
Definition pair1_simpl : forall (d:D) (x:unit), pair1 d x = d.
|
||||
destruct x; trivial.
|
||||
Defined.
|
||||
|
||||
Definition Pair1 : D -m> Dprodi DI1.
|
||||
exists pair1; red; intros; intro d.
|
||||
repeat (rewrite pair1_simpl);trivial.
|
||||
Defined.
|
||||
|
||||
|
||||
Lemma Pair1_simpl : forall (d:D), Pair1 d = pair1 d.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Definition PAIR1 : D -c> Dprodi DI1.
|
||||
exists Pair1; red; intros; repeat (rewrite Pair1_simpl).
|
||||
intro d; rewrite pair1_simpl.
|
||||
rewrite (Dprodi_lub_simpl (Di:=DI1)).
|
||||
apply lub_le_compat; intros.
|
||||
intro x; simpl; rewrite pair1_simpl; auto.
|
||||
Defined.
|
||||
|
||||
Lemma pi_simpl : forall (d:D), pi (pair1 d) = d.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Definition DI1_map (f : D -c> D)
|
||||
: Dprodi DI1 -c> Dprodi DI1 :=
|
||||
DMAPi (fun t:unit => f).
|
||||
|
||||
Lemma DI1_map_eq : forall (f : D -c> D) (d:Dprodi DI1),
|
||||
DI1_map f d == pair1 (f (pi d)).
|
||||
intros; simpl; apply Oprodi_eq_intro; intro b; case b; trivial.
|
||||
Save.
|
||||
End Product1.
|
||||
|
||||
Hint Resolve DI1_map_eq.
|
||||
End of Stuff *)
|
291
papers/domains-2009/PredomSum.v
Executable file
291
papers/domains-2009/PredomSum.v
Executable file
|
@ -0,0 +1,291 @@
|
|||
(*==========================================================================
|
||||
Definition of co-products and associated lemmas
|
||||
==========================================================================*)
|
||||
Require Import PredomCore.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Order-theoretic definitions
|
||||
==========================================================================*)
|
||||
Definition Osum: ord -> ord -> ord.
|
||||
intros O1 O2. exists ((O1 + O2)%type)
|
||||
(fun x y => match x, y with | inl x, inl y => x <= y | inr x, inr y => x <= y | _,_ => False end).
|
||||
intros x. case x ; intros t ; apply Ole_refl.
|
||||
intros x y z. case x ; clear x ; intros x.
|
||||
case y ; clear y ; intros y ; try (intros ; intuition ; fail).
|
||||
case z ; clear z ; intros z ; try (intros ; intuition ; fail). apply Ole_trans.
|
||||
case y ; clear y ; intros y ; try (intros ; intuition ; fail).
|
||||
case z ; clear z ; intros z ; try (intros ; intuition ; fail). apply Ole_trans.
|
||||
Defined.
|
||||
|
||||
Definition Inl (O1 O2:ord) : O1 -m> (Osum O1 O2).
|
||||
intros. exists (fun x => @inl O1 O2 x). unfold monotonic.
|
||||
intros. auto.
|
||||
Defined.
|
||||
|
||||
Definition Inr (O1 O2 :ord) : O2 -m> (Osum O1 O2).
|
||||
intros. exists (fun x => @inr O1 O2 x). unfold monotonic. auto.
|
||||
Defined.
|
||||
|
||||
Definition sum_left (D1 D2:cpo) (c:natO -m> Osum D1 D2) (ex:{d| c O = inl D2 d}) : (natO -m> D1).
|
||||
intros D1 D2 c ex.
|
||||
destruct ex as [d ex].
|
||||
exists (fun x => match c x with | inl y => y | inr _ => d end).
|
||||
unfold monotonic.
|
||||
intros n m lm. assert (c n <= c m). auto. generalize H. clear H.
|
||||
assert ((O <= n) % nat) as no by omega. assert (c O <= c n) as con by auto. rewrite ex in con. clear ex no.
|
||||
generalize con. clear con.
|
||||
case (c n). intros d1 _. case (c m). intros d1'. auto.
|
||||
intros d2' incon. inversion incon. intros d1' incon. inversion incon.
|
||||
Defined.
|
||||
|
||||
Definition sum_right (D1 D2:cpo) (c:natO -m> Osum D1 D2) (ex:{d | c O = inr D1 d}) : (natO -m> D2).
|
||||
intros D1 D2 c ex.
|
||||
destruct ex as [d ex].
|
||||
exists (fun x => match c x with | inr y => y | inl _ => d end).
|
||||
unfold monotonic.
|
||||
intros n m lm. assert (c n <= c m). auto. generalize H. clear H.
|
||||
assert ((O <= n) % nat) as no by omega. assert (c O <= c n) as con by auto. rewrite ex in con. clear ex no.
|
||||
generalize con. clear con.
|
||||
case (c n). intros d1 _. case (c m). intros d1'. auto.
|
||||
intros d2' incon. inversion incon. intros d1' _. case (c m). intros tt incon. inversion incon.
|
||||
intros d2' l. auto.
|
||||
Defined.
|
||||
|
||||
Definition SuminjProof: forall (D1 D2:cpo) (c:natO -m> Osum D1 D2) (n:nat),
|
||||
{d | c n = inl D2 d} + {d | c n = inr D1 d}.
|
||||
intros D1 D2 c n.
|
||||
assert ((0 <= n)%nat). omega.
|
||||
assert (c 0 <= c n). auto. clear H.
|
||||
case_eq (c 0).
|
||||
intros d1 c0.
|
||||
case_eq (c n). intros dn cnd.
|
||||
left.
|
||||
exists dn. auto.
|
||||
intros d2 incon. rewrite incon in H0. rewrite c0 in H0. inversion H0.
|
||||
|
||||
intros d2 c0. right.
|
||||
case_eq (c n). intros dn cnd.
|
||||
rewrite cnd in H0. rewrite c0 in H0. inversion H0.
|
||||
intros dn cn. exists dn. auto.
|
||||
Defined.
|
||||
|
||||
Definition sum_lub (D1 D2:cpo) (c: natO -m> Osum D1 D2) : (Osum D1 D2).
|
||||
intros D1 D2 c. case (SuminjProof c 0).
|
||||
intros X. refine (inl D2 (lub (sum_left X))).
|
||||
intro X. refine (inr D1 (lub (sum_right X))).
|
||||
Defined.
|
||||
|
||||
(*==========================================================================
|
||||
Domain-theoretic definitions
|
||||
==========================================================================*)
|
||||
Definition Dsum : cpo -> cpo -> cpo.
|
||||
intros D1 D2. exists (Osum D1 D2) (@sum_lub D1 D2).
|
||||
intros c n.
|
||||
case_eq (c 0).
|
||||
intros d1 c0.
|
||||
simpl. case_eq (c n). intros dn cn.
|
||||
unfold sum_lub. destruct (SuminjProof c 0) as [inj|inj].
|
||||
assert (sum_left inj n = dn).
|
||||
destruct inj. simpl. rewrite cn. auto.
|
||||
rewrite <- H.
|
||||
apply le_lub.
|
||||
|
||||
destruct inj as [dd c0n]. rewrite c0n in c0. inversion c0.
|
||||
intros d2 cn.
|
||||
assert ((0 <= n) % nat) as A by omega.
|
||||
assert (c 0 <= c n) by auto. rewrite c0 in H. rewrite cn in H. inversion H.
|
||||
|
||||
intros d2 c0. simpl. case_eq (c n). intros dn cn.
|
||||
unfold sum_lub. destruct (SuminjProof c 0) as [[d inj] | inj].
|
||||
rewrite inj in c0. inversion c0.
|
||||
assert ((0 <= n) % nat) as A by omega. assert (c 0 <= c n) by auto.
|
||||
rewrite c0 in H. rewrite cn in H. inversion H.
|
||||
intros dn cn. unfold sum_lub. destruct (SuminjProof c 0) as [[dd inj]|[d inj]].
|
||||
rewrite c0 in inj. inversion inj.
|
||||
assert (sum_right (exist (fun dd => c 0 = inr D1 dd) d inj) n = dn).
|
||||
simpl. rewrite cn. auto.
|
||||
rewrite <- H.
|
||||
apply (le_lub).
|
||||
|
||||
intros c x. case x ; clear x ; intros x cn. unfold sum_lub.
|
||||
case (SuminjProof c 0). intros s.
|
||||
assert (lub (sum_left s) <= x).
|
||||
apply (lub_le). intros n. specialize (cn n). destruct s as [d c0].
|
||||
simpl. case_eq (c n). intros dn dnl. rewrite dnl in cn. auto.
|
||||
intros dn cnr. rewrite cnr in cn. inversion cn. auto.
|
||||
intros s. destruct s. simpl. specialize (cn 0).
|
||||
rewrite e in cn. inversion cn.
|
||||
|
||||
unfold sum_lub. case (SuminjProof c 0). intros s.
|
||||
destruct s. simpl. specialize (cn 0). rewrite e in cn. inversion cn.
|
||||
intros s. assert (lub (sum_right s) <= x).
|
||||
apply lub_le. intros n. specialize (cn n). destruct s as [d c0].
|
||||
simpl. case_eq (c n). intros dl cnl. rewrite cnl in cn. inversion cn.
|
||||
intros dl dll. rewrite dll in cn. auto. auto.
|
||||
Defined.
|
||||
|
||||
Definition Sum_fun : forall (D1 D2 D3:cpo) (f:D1 -m> D3) (g:D2-m> D3),
|
||||
Dsum D1 D2 -m> D3.
|
||||
intros. exists (fun p => match p with inl y => (f y) | inr y => (g y) end).
|
||||
unfold monotonic. intros x y.
|
||||
case_eq x. intros xx _. clear x. case_eq y. intros yy _ ll. auto.
|
||||
intros yy _ incon. inversion incon.
|
||||
intros xx _. case_eq y. intros yy _ incon. inversion incon.
|
||||
intros yy _ ll. auto.
|
||||
Defined.
|
||||
|
||||
Definition SUM_fun (D1 D2 D3:cpo)(f:D1 -c> D3) (g:D2 -c> D3) :
|
||||
Dsum D1 D2 -c> D3.
|
||||
intros. exists (Sum_fun (fcontit f) (fcontit g)).
|
||||
unfold continuous. intros c.
|
||||
simpl. case_eq (sum_lub c).
|
||||
intros ll. intros sl. unfold sum_lub in sl.
|
||||
destruct (SuminjProof c 0).
|
||||
inversion sl. clear sl.
|
||||
destruct s as [dd c0]. assert (continuous (fcontit f)) as fC by auto.
|
||||
unfold continuous in fC. rewrite fC.
|
||||
apply lub_le_compat. simpl. intros n.
|
||||
case_eq (c n). intros t cn. simpl. simpl in cn. rewrite cn. auto.
|
||||
intros t cn. simpl. simpl in cn. rewrite cn.
|
||||
assert (c 0 <= c n). assert (0 <= n) % nat by omega. auto.
|
||||
simpl in c0. simpl in H. rewrite c0 in H. rewrite cn in H. inversion H.
|
||||
inversion sl.
|
||||
|
||||
intros ll. intros sl. unfold sum_lub in sl. destruct (SuminjProof c 0).
|
||||
inversion sl. inversion sl. clear sl. destruct s as [dd c0].
|
||||
assert (continuous (fcontit g)) as gC by auto.
|
||||
unfold continuous in gC. rewrite gC. apply lub_le_compat.
|
||||
simpl. intros n.
|
||||
simpl. case_eq (c n). intros t cn. simpl. simpl in cn. rewrite cn.
|
||||
assert (c 0 <= c n). assert (0 <= n)%nat as A by omega. auto.
|
||||
simpl in H. rewrite c0 in H. rewrite cn in H. inversion H.
|
||||
intros t cn. simpl in cn. rewrite cn. auto.
|
||||
Defined.
|
||||
|
||||
Implicit Arguments SUM_fun [D1 D2 D3].
|
||||
|
||||
Add Parametric Morphism (D E F:cpo) : (@SUM_fun D E F)
|
||||
with signature (@Ole (D -c> F)) ++> (@Ole (E -c> F)) ++> (@Ole ((Dsum D E) -c> F))
|
||||
as SUM_fun_le_compat.
|
||||
intros f f' lf g g' lg.
|
||||
simpl. intros x ; case x ; clear x ; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E F:cpo) : (@SUM_fun D E F)
|
||||
with signature (@Oeq (D -c> F)) ==> (@Oeq (E -c> F)) ==> (@Oeq ((Dsum D E) -c> F))
|
||||
as SUM_fun_eq_compat.
|
||||
intros f f' lf g g' lg.
|
||||
simpl. apply fcont_eq_intro.
|
||||
destruct lg. destruct lf.
|
||||
intros x ; split ; case x ; clear x ; auto.
|
||||
Qed.
|
||||
|
||||
Definition INL (D1 D2:cpo) : D1 -c> Dsum D1 D2.
|
||||
intros. exists (@Inl D1 D2). unfold continuous.
|
||||
intros c. simpl. apply lub_le_compat.
|
||||
simpl. auto.
|
||||
Defined.
|
||||
|
||||
Implicit Arguments INL [D1 D2].
|
||||
|
||||
Add Parametric Morphism (D E:cpo) : (@INL D E)
|
||||
with signature (@Ole D) ++> (@Ole (Dsum D E))
|
||||
as SUM_INL_le_compat.
|
||||
intros ; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E:cpo) : (@INL D E)
|
||||
with signature (@Oeq D) ==> (@Oeq (Dsum D E))
|
||||
as SUM_INL_eq_compat.
|
||||
intros ; auto.
|
||||
Qed.
|
||||
|
||||
Definition INR (D1 D2:cpo) : D2 -c> Dsum D1 D2.
|
||||
intros. exists (@Inr D1 D2). unfold continuous. intros c. simpl. apply lub_le_compat.
|
||||
auto.
|
||||
Defined.
|
||||
|
||||
Implicit Arguments INR [D1 D2].
|
||||
|
||||
Add Parametric Morphism (D E:cpo) : (@INR D E)
|
||||
with signature (@Ole E) ++> (@Ole (Dsum D E))
|
||||
as SUM_INR_le_compat.
|
||||
intros ; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E:cpo) : (@INR D E)
|
||||
with signature (@Oeq E) ==> (@Oeq (Dsum D E))
|
||||
as SUM_INR_eq_compat.
|
||||
intros ; auto.
|
||||
Qed.
|
||||
|
||||
Definition SUM_map D1 D2 D3 D4 (f:D1 -c> D3) (g:D2 -c> D4) :=
|
||||
SUM_fun (INL << f) (INR << g).
|
||||
|
||||
Add Parametric Morphism (D E F G:cpo) : (@SUM_map D E F G)
|
||||
with signature (@Ole (D -c> F)) ++> (@Ole (E -c> G)) ++> (@Ole ((Dsum D E) -c> (Dsum F G)))
|
||||
as SUM_map_le_compat.
|
||||
intros f f' lf g g' lg.
|
||||
simpl. intros x ; case x ; clear x ; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E F G:cpo) : (@SUM_map D E F G)
|
||||
with signature (@Oeq (D -c> F)) ==> (@Oeq (E -c> G)) ==> (@Oeq ((Dsum D E) -c> (Dsum F G)))
|
||||
as SUM_map_eq_compat.
|
||||
intros f f' lf g g' lg.
|
||||
simpl. apply fcont_eq_intro.
|
||||
destruct lg. destruct lf.
|
||||
intros x ; split ; case x ; clear x ; auto.
|
||||
Qed.
|
||||
|
||||
Definition SUM_Map1 (D E F G:cpo) : (D -C-> F) -> (E -C-> G) -M-> Dsum D E -C-> Dsum F G.
|
||||
intros. exists (@SUM_map D E F G X).
|
||||
unfold monotonic. intros x y xyl. intros xx ; case xx ; auto.
|
||||
Defined.
|
||||
|
||||
Definition SUM_Map (D E F G:cpo) : (D -C-> F) -M-> (E -C-> G) -M-> Dsum D E -C-> Dsum F G.
|
||||
intros. exists (@SUM_Map1 D E F G). unfold monotonic.
|
||||
intros x y xyl. intros xx yy ; case xx ; case yy ; simpl ; auto.
|
||||
Defined.
|
||||
|
||||
Lemma SUM_Map_continuous2 : forall D E F G, continuous2 (SUM_Map D E F G).
|
||||
intros. unfold continuous2. intros c1 c2.
|
||||
intros p. case p ; auto.
|
||||
Save.
|
||||
|
||||
Definition SUM_MAP (D E F G:cpo) : (D -C-> F) -C-> (E -C-> G) -C-> Dsum D E -C-> Dsum F G :=
|
||||
(continuous2_cont (@SUM_Map_continuous2 D E F G)).
|
||||
|
||||
Lemma SUM_map_simpl (D E F G:cpo) f g : SUM_MAP D E F G f g = SUM_map f g.
|
||||
intros. auto.
|
||||
Qed.
|
||||
|
||||
Lemma SUM_fun_simpl (D E F:cpo) f g s : @SUM_fun D E F f g s = Sum_fun (fcontit f) (fcontit g) s.
|
||||
intros. auto.
|
||||
Qed.
|
||||
|
||||
Lemma Dsum_lcom: forall D1 D2 D3 f g, @SUM_fun D1 D2 D3 f g << INL == f.
|
||||
intros. apply fcont_eq_intro. intros s. auto.
|
||||
Qed.
|
||||
|
||||
Lemma Dsum_rcom: forall D1 D2 D3 f g, @SUM_fun D1 D2 D3 f g << INR == g.
|
||||
intros. apply fcont_eq_intro. auto.
|
||||
Qed.
|
||||
|
||||
Lemma Dsum_unique: forall D1 D2 D3 f g h, h << INL == f -> h << INR == g ->
|
||||
h == @SUM_fun D1 D2 D3 f g.
|
||||
intros. apply fcont_eq_intro. intros x.
|
||||
case x. clear x. intros t.
|
||||
assert (xx:= fcont_eq_elim H t). auto.
|
||||
intros t. assert (xx:= fcont_eq_elim H0 t). auto.
|
||||
Qed.
|
||||
|
||||
Lemma SUM_fun_comp :
|
||||
forall C D E F (f:C-c>E) (g:D-c>E) (h:E-c>F), h << SUM_fun f g == SUM_fun (h << f) (h << g).
|
||||
intros C D E F f g h.
|
||||
apply fcont_eq_intro. intro x. case x; auto.
|
||||
Qed.
|
||||
|
198
papers/domains-2009/Sets.v
Executable file
198
papers/domains-2009/Sets.v
Executable file
|
@ -0,0 +1,198 @@
|
|||
Require Export Setoids.Setoid.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
Record SetKind :=
|
||||
{
|
||||
tSetKind :> Type;
|
||||
eq_rel : tSetKind -> tSetKind -> Prop;
|
||||
Peq_rel : equiv _ eq_rel
|
||||
}.
|
||||
|
||||
Record set (sk:SetKind) :=
|
||||
{
|
||||
car : tSetKind sk -> Prop;
|
||||
car_eq : forall x y, @eq_rel sk x y -> car x -> car y
|
||||
}.
|
||||
|
||||
Lemma set_trans : forall (T:SetKind) a b c, (@eq_rel T a b) -> (eq_rel b c) -> (eq_rel a c).
|
||||
intros T a b c ab bc.
|
||||
assert (xx:=Peq_rel T). unfold equiv in *. destruct xx as [_ [tt _]].
|
||||
unfold transitive in tt. apply (tt _ _ _ ab bc).
|
||||
Qed.
|
||||
|
||||
Hint Resolve set_trans.
|
||||
|
||||
Lemma set_refl : forall (T:SetKind) a, @eq_rel T a a.
|
||||
intros T a. assert (xx:=Peq_rel T). unfold equiv in *. destruct xx as [tt [_ _]].
|
||||
unfold reflexive in tt. apply (tt a).
|
||||
Qed.
|
||||
|
||||
Hint Resolve set_refl.
|
||||
|
||||
Lemma set_sym : forall (T:SetKind) a b, (@eq_rel T a b) -> (eq_rel b a).
|
||||
intros T a b. assert (xx:=Peq_rel T). unfold equiv in *. destruct xx as [_ [_ tt]].
|
||||
unfold symmetric in tt. apply (tt a b).
|
||||
Qed.
|
||||
|
||||
Hint Resolve set_sym.
|
||||
|
||||
Definition union T (S R:set T) : set T.
|
||||
intros T S R. exists (fun x => car S x \/ car R x).
|
||||
intros x y Rel xy. destruct xy as [Sx | Sy].
|
||||
left. apply (car_eq Rel Sx). right. apply (car_eq Rel Sy).
|
||||
Defined.
|
||||
|
||||
Definition Union T (S:set T -> Prop) : set T.
|
||||
intros T S. exists (fun x => exists s, S s /\ car s x).
|
||||
intros x y Rel xy. destruct xy as [s [Ss cx]].
|
||||
exists s. split. auto. refine (car_eq Rel cx).
|
||||
Defined.
|
||||
|
||||
Definition intersect T (S R:set T) : set T.
|
||||
intros T S R. exists (fun x => car S x /\ car R x).
|
||||
intros x y Rel [Sx Rx]. split.
|
||||
apply (car_eq Rel Sx). apply (car_eq Rel Rx).
|
||||
Defined.
|
||||
|
||||
Definition Intersect T (S: set T -> Prop) : set T.
|
||||
intros T S. exists (fun x => forall s, S s -> car s x).
|
||||
intros x y Rel C s Ss.
|
||||
apply (car_eq Rel (C _ Ss)).
|
||||
Defined.
|
||||
|
||||
Definition subset T (S R:set T) := forall x, car S x -> car R x.
|
||||
|
||||
Definition set_eq T (S R:set T) := subset S R /\ subset R S.
|
||||
|
||||
Definition setmorphism (T T':SetKind) (f: T -> set T') :=
|
||||
forall x y, @eq_rel T x y -> set_eq (f x) (f y).
|
||||
|
||||
Record fsetm (T T':SetKind) :=
|
||||
{
|
||||
fsett :> T -> set T';
|
||||
fsetmorphism : setmorphism fsett
|
||||
}.
|
||||
|
||||
Definition Sarrow (T:Type) (T':SetKind) : SetKind.
|
||||
intros T T'. exists (T -> T') (fun f g => forall x, eq_rel (f x) (g x)).
|
||||
assert (equiv _ (@eq_rel T')). apply Peq_rel.
|
||||
unfold equiv in *. unfold reflexive in *. unfold transitive in *. unfold symmetric in *.
|
||||
destruct H as [r [t s]]. split.
|
||||
intros f x. auto. split. intros f g h a b x. specialize (a x). specialize (b x).
|
||||
specialize (t _ _ _ a b). apply t.
|
||||
intros f g a x. specialize (a x). specialize (s _ _ a). apply s.
|
||||
Defined.
|
||||
|
||||
Definition Sprod (T T' : SetKind) : SetKind.
|
||||
intros T T'. exists (T * T')%type (fun p0 p1 => @eq_rel T (fst p0) (fst p1) /\ @eq_rel T' (snd p0) (snd p1)).
|
||||
assert (equiv _ (@eq_rel T)) as [sr [st ss]] by (apply Peq_rel).
|
||||
assert (equiv _ (@eq_rel T')) as [s'r [s't s's]] by (apply Peq_rel).
|
||||
unfold equiv. unfold reflexive in *. unfold transitive in *. unfold symmetric in *.
|
||||
split. intros x. case x. clear x. intros ; auto. split.
|
||||
intros x ; case x ; clear x. intros x0 x1 y ; case y ; clear y.
|
||||
intros y0 y1 z ; case z ; clear z ; intro z0. simpl. intros z1 [E0 E1] [E2 E3].
|
||||
specialize (st _ _ _ E0 E2). specialize (s't _ _ _ E1 E3).
|
||||
split ; auto.
|
||||
intros x ; case x ; clear x ; intros x0 x1 y ; case y ; clear y ; intros y0 y1.
|
||||
simpl. intros [E0 E1]. specialize (ss _ _ E0). specialize (s's _ _ E1).
|
||||
split ; auto.
|
||||
Defined.
|
||||
|
||||
Definition SFst T T' : set (Sprod T T') -> set T.
|
||||
intros T T' S. exists (fun t => exists t', car S (t,t')).
|
||||
intros t0 t1. intros Rel ex. destruct ex as [t' c0]. exists t'.
|
||||
simpl in *. assert (@eq_rel (Sprod T T') (t0,t') (t1,t')) as Rel'.
|
||||
simpl. split. apply Rel. apply set_refl.
|
||||
refine (car_eq Rel' c0).
|
||||
Defined.
|
||||
|
||||
Definition SSnd T T' : set (Sprod T T') -> set T'.
|
||||
intros T T' S. exists (fun t' => exists t, car S (t,t')).
|
||||
intros t0 t1. intros Rel ex. destruct ex as [t' c0]. exists t'.
|
||||
simpl in *. assert (@eq_rel (Sprod T T') (t',t0) (t',t1)) as Rel'.
|
||||
simpl. split. apply set_refl. apply Rel.
|
||||
refine (car_eq Rel' c0).
|
||||
Defined.
|
||||
|
||||
Definition Ssum (T T':SetKind) : SetKind.
|
||||
intros T T'. exists (T + T') % type
|
||||
(fun x y => match x, y with | inl x, inl y => eq_rel x y
|
||||
| inr x, inr y => eq_rel x y | _,_ => False
|
||||
end).
|
||||
unfold equiv. split. unfold reflexive. intros x ; case x ; auto.
|
||||
split. unfold transitive.
|
||||
intros x y z ; case x ; clear x ; intros x ; case y ; clear y ; intros y ;
|
||||
case z ; clear z ; intros z ; auto.
|
||||
refine (@set_trans T x y z).
|
||||
intros incon. inversion incon. intros incon. inversion incon.
|
||||
refine (@set_trans T' x y z).
|
||||
|
||||
unfold symmetric.
|
||||
intros x y ; case x ; clear x ; intro x ; case y ; clear y ; intros y ; auto.
|
||||
Defined.
|
||||
|
||||
Definition SingleSet (T:SetKind) (v:T) : set T.
|
||||
intros T x. exists (fun (y:T) => eq_rel x y).
|
||||
intros x0 y0. intros veq xeq. assert (tt:= proj1 (proj2 (Peq_rel T))).
|
||||
unfold transitive in tt. apply (tt _ _ _ xeq veq).
|
||||
Defined.
|
||||
|
||||
Definition EmptySet (T:SetKind) : set T.
|
||||
intros T. exists (fun (x:T) => False).
|
||||
auto.
|
||||
Defined.
|
||||
|
||||
Definition inv_Kind (T:SetKind) S (f:S -> T) : SetKind.
|
||||
intros T S f. exists (S) (fun x y => eq_rel (f x) (f y)).
|
||||
unfold equiv. split. unfold reflexive.
|
||||
intros s. auto. split. unfold transitive. intros x y z xy yz.
|
||||
apply (proj1 (proj2 (Peq_rel T)) (f x) (f y) (f z)) ; auto.
|
||||
intros x y. apply (proj2 (proj2 (Peq_rel T)) (f x) (f y)).
|
||||
Defined.
|
||||
|
||||
Definition inv_image (T:SetKind) S (f:S -> T) : set T -> set (inv_Kind f).
|
||||
intros T S f t.
|
||||
exists (fun x => car t (f x)).
|
||||
simpl. intros x y xy cx.
|
||||
refine (car_eq xy cx).
|
||||
Defined.
|
||||
|
||||
Definition LeibnizKind (T:Type) : SetKind.
|
||||
intros T. exists T (@eq T).
|
||||
unfold equiv. split. unfold reflexive. auto. split.
|
||||
unfold transitive. intros x y z xy yz. apply (trans_eq xy yz).
|
||||
unfold symmetric. intros x y xy. apply (sym_eq xy).
|
||||
Defined.
|
||||
|
||||
Definition LeibnizSet (T:Type) (car:T -> Prop) : set (LeibnizKind T).
|
||||
intros T c. exists c.
|
||||
intros x y eq. simpl in eq. rewrite eq. auto.
|
||||
Defined.
|
||||
|
||||
Definition EqSetKind T eq_rel (Peq_rel:equiv T eq_rel) : SetKind.
|
||||
intros T eq equiv. exists T eq. auto.
|
||||
Defined.
|
||||
|
||||
Definition UNIV T : set T.
|
||||
intros T. exists (fun x => True). auto.
|
||||
Defined.
|
||||
|
||||
Definition set_respect (T T':SetKind) (f:T -> T') := forall x y, eq_rel x y -> eq_rel (f x) (f y).
|
||||
|
||||
Definition Inv_image (T T':SetKind) (f:T -> T') (sf:set_respect f) : (set T') -> set T.
|
||||
intros T T' f sf s. exists (fun x => car s (f x)).
|
||||
intros x y eq c. refine (car_eq _ c). unfold set_respect in sf.
|
||||
refine (sf _ _ eq).
|
||||
Defined.
|
||||
|
||||
Lemma fst_set_respect : forall (T T':SetKind), @set_respect (Sprod T T') T (@fst T T').
|
||||
intros T T'. unfold set_respect.
|
||||
simpl. intros x y [R _]. apply R.
|
||||
Qed.
|
||||
|
||||
Lemma snd_set_respect : forall (T T':SetKind), @set_respect (Sprod T T') T' (@snd T T').
|
||||
intros T T' x y [_ R]. apply R.
|
||||
Qed.
|
||||
|
1159
papers/domains-2009/lc.v
Executable file
1159
papers/domains-2009/lc.v
Executable file
File diff suppressed because it is too large
Load diff
453
papers/domains-2009/typedadequacy.v
Executable file
453
papers/domains-2009/typedadequacy.v
Executable file
|
@ -0,0 +1,453 @@
|
|||
Require Import utility.
|
||||
Require Import PredomAll.
|
||||
Require Import typedlambda.
|
||||
Require Import typedopsem.
|
||||
Require Import typeddensem.
|
||||
|
||||
Set Implicit Arguments.
|
||||
|
||||
(*==========================================================================
|
||||
Logical relation between semantics and terms, used to prove adequacy
|
||||
See Winskel Chapter 11, Section 11.4.
|
||||
==========================================================================*)
|
||||
|
||||
(** printing senv %\ensuremath{\rho}% *)
|
||||
(** printing d1 %\ensuremath{d_1}% *)
|
||||
|
||||
(* Lift a value-relation to a computation-relation *)
|
||||
Definition liftRel ty (R : SemTy ty -> CValue ty -> Prop) :=
|
||||
fun (d:DL _) e => forall d', d == Val d' -> exists v, e =>> v /\ R d' v.
|
||||
|
||||
(* Logical relation between semantics and closed value terms *)
|
||||
Unset Implicit Arguments.
|
||||
|
||||
Fixpoint relVal ty : SemTy ty -> CValue ty -> Prop :=
|
||||
match ty with
|
||||
| Int => fun d v => v = TINT d
|
||||
| Bool => fun d v => v = TBOOL d
|
||||
| ty1 --> ty2 => fun d v => exists e, v = TFIX e /\ forall d1 v1, relVal ty1 d1 v1 -> liftRel (relVal ty2) (d d1) (substExp [ v1, v ] e)
|
||||
| ty1 ** ty2 => fun d v => exists v1, exists v2, v = TPAIR v1 v2 /\ relVal ty1 (FST d) v1 /\ relVal ty2 (SND d) v2
|
||||
end.
|
||||
|
||||
(* Logical relation between semantic environments and value environments *)
|
||||
Fixpoint relEnv env : SemEnv env -> Subst env nil -> Prop :=
|
||||
match env with
|
||||
| nil => fun _ _ => True
|
||||
| ty :: env => fun d s => relVal ty (SND d) (hdMap s) /\ relEnv env (FST d) (tlMap s)
|
||||
end.
|
||||
|
||||
(* Computation relation, already inlined into relVal *)
|
||||
Definition relExp ty := liftRel (relVal ty).
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
|
||||
|
||||
Lemma relVal_lower_lift :
|
||||
forall ty,
|
||||
(forall d d' (v:CValue ty), d <= d' -> relVal ty d' v -> relVal ty d v) ->
|
||||
(forall d d' (e:CExp ty), d <= d' -> relExp ty d' e -> relExp ty d e).
|
||||
intros t IH d d' e l R.
|
||||
unfold relExp, liftRel in *.
|
||||
intros d'0 eq.
|
||||
rewrite eq in l.
|
||||
assert (S := DLle_Val_exists_eq l).
|
||||
destruct S as [d'1 [d'eq l']].
|
||||
clear eq l. specialize (R d'1 d'eq).
|
||||
destruct R as [v [ev grl]].
|
||||
exists v. split; auto. specialize (IH d'0 d'1 v l'). auto.
|
||||
Qed.
|
||||
|
||||
(* Lemma 11.12 (ii) from Winskel *)
|
||||
Lemma relVal_lower: forall ty d d' v, d <= d' -> relVal ty d' v -> relVal ty d v.
|
||||
induction ty.
|
||||
|
||||
(* Int *)
|
||||
simpl. intros. subst. auto.
|
||||
|
||||
(* Bool *)
|
||||
simpl. intros. subst. auto.
|
||||
|
||||
(* Arrow *)
|
||||
intros d d' v l. simpl relVal. intros [b [veq C]]. subst. exists b.
|
||||
split; auto.
|
||||
intros d1 v1 gl1.
|
||||
assert (S := fcont_le_elim l d1).
|
||||
specialize (C d1 v1 gl1). unfold liftRel in *.
|
||||
assert (A := relVal_lower_lift IHty2 S (e := (substExp (env':= nil) [v1,TFIX b] b))). unfold relExp, liftRel in A. auto.
|
||||
|
||||
(* Pair *)
|
||||
simpl. intros p1 p2 v [leq1 leq2] [v1 [v2 [veq [gl1 gl2]]]].
|
||||
subst.
|
||||
exists v1. exists v2. split; auto.
|
||||
specialize (IHty1 _ _ v1 leq1 gl1).
|
||||
specialize (IHty2 _ _ v2 leq2 gl2). split ; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism ty (v:CValue ty) : (fun d => relVal ty d v)
|
||||
with signature (@Oeq (SemTy ty)) ==> iff as relVal_eq_compat.
|
||||
intros x y xyeq.
|
||||
destruct xyeq as [xy1 xy2].
|
||||
split; apply relVal_lower; trivial.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism ty (v:CExp ty) : (fun d => relExp ty d v)
|
||||
with signature (@Oeq (DL (SemTy ty))) ==> iff as relExp_eq_compat.
|
||||
intros x y xyeq.
|
||||
destruct xyeq as [xy1 xy2].
|
||||
split; apply (relVal_lower_lift (relVal_lower (ty:=ty))); trivial.
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Admissibility
|
||||
==========================================================================*)
|
||||
|
||||
Lemma rel_admissible_lift : forall ty,
|
||||
(forall v, admissible (fun d => relVal ty d v)) ->
|
||||
(forall e, admissible (fun d => relExp ty d e)).
|
||||
intros t IH e.
|
||||
unfold admissible in *.
|
||||
intros c C.
|
||||
unfold relExp, liftRel in *.
|
||||
intros d eq.
|
||||
destruct (lubval eq) as [k [dv [ck dvd]]].
|
||||
destruct (C _ _ ck) as [v [ev rv]].
|
||||
exists v. split ; auto.
|
||||
destruct (DLvalgetchain ck) as [c' P].
|
||||
specialize (IH v c').
|
||||
assert ((lub (c:=SemTy t) c') == d) as deq. apply vinj.
|
||||
rewrite <- eq. rewrite <- eta_val. apply (Oeq_trans (y:=lub (fcontit eta @ c'))).
|
||||
apply lub_comp_eq ; auto.
|
||||
apply Oeq_sym.
|
||||
refine (Oeq_trans (lub_lift_left _ k) _).
|
||||
apply lub_eq_compat. apply fmon_eq_intro. intros x. simpl. apply P.
|
||||
|
||||
refine (relVal_lower (proj2 deq) _).
|
||||
apply IH. intros n.
|
||||
specialize (P n). specialize (C _ _ P).
|
||||
destruct C as [vv [evv rvv]].
|
||||
rewrite (Determinacy ev evv). apply rvv.
|
||||
Qed.
|
||||
|
||||
Require Import Program.Equality.
|
||||
|
||||
(* Lemma 11.12 (iii) from Winskel *)
|
||||
Lemma rel_admissible: forall ty v, admissible (fun d => relVal ty d v).
|
||||
induction ty.
|
||||
|
||||
(* Int *)
|
||||
intros v c C. simpl. apply (C 0).
|
||||
|
||||
(* Bool *)
|
||||
intros v c C. simpl. apply (C 0).
|
||||
|
||||
(* Arrow *)
|
||||
unfold admissible in *.
|
||||
intros v c C.
|
||||
simpl relVal.
|
||||
assert (rv := C 0). simpl in rv. destruct rv as [b [tbeq C0]].
|
||||
exists b.
|
||||
split; auto.
|
||||
intros d1 v1 R1.
|
||||
refine (rel_admissible_lift _ _). unfold admissible.
|
||||
intros vv cc CC. apply IHty2. apply CC.
|
||||
|
||||
intros n. unfold relExp. unfold liftRel.
|
||||
intros d' CC.
|
||||
assert (Cn := C n). simpl in Cn.
|
||||
destruct Cn as [bb [tteq Cn]].
|
||||
assert (b = bb) as bbeq. rewrite tteq in *.
|
||||
clear C0 d1 v1 R1 CC tteq v Cn C.
|
||||
generalize bb tbeq. clear bb tbeq.
|
||||
intros bb. intros eq.
|
||||
apply (TFIX_injective (sym_equal eq)).
|
||||
rewrite <- bbeq in *. clear bb bbeq.
|
||||
clear tteq.
|
||||
specialize (Cn _ _ R1).
|
||||
unfold liftRel in Cn. specialize (Cn _ CC). apply Cn.
|
||||
|
||||
(* Pair *)
|
||||
unfold admissible.
|
||||
intros v c C. simpl. unfold prod_lub.
|
||||
generalize C.
|
||||
destruct (ClosedPair v) as [v1 [v2 veq]]. subst.
|
||||
intros _.
|
||||
exists v1. exists v2.
|
||||
split;auto.
|
||||
specialize (IHty1 v1). specialize (IHty2 v2).
|
||||
split. rewrite FST_simpl. rewrite Fst_simpl. simpl fst. apply IHty1. intros n. specialize (C n). simpl. simpl in C.
|
||||
destruct C as [v3 [v4 [veq [C1 C2]]]]. clear C2.
|
||||
destruct (TPAIR_injective veq) as [E1 E2]. subst.
|
||||
trivial.
|
||||
|
||||
rewrite SND_simpl. rewrite Snd_simpl. simpl snd. apply IHty2. intros n. specialize (C n). simpl. simpl in C.
|
||||
destruct C as [v3 [v4 [veq [C1 C2]]]]. clear C1.
|
||||
destruct (TPAIR_injective veq) as [E1 E2]. subst.
|
||||
trivial.
|
||||
|
||||
Qed.
|
||||
|
||||
Lemma PBot_val_incon: forall D x, DL_bot D == Val (D:=D) x -> False.
|
||||
intros D x incon. assert (xx := proj2 incon). clear incon.
|
||||
simpl in xx. inversion xx. subst. clear H1 xx. induction n.
|
||||
simpl in H0. rewrite DL_bot_eq in H0. inversion H0.
|
||||
rewrite DL_bot_eq in H0. simpl in H0. apply (IHn H0).
|
||||
Qed.
|
||||
|
||||
Lemma choose_t_simpl: forall (D:cpo) e1 e2, choose D e1 e2 true = e1.
|
||||
intros D e1 e2. unfold choose. auto.
|
||||
Qed.
|
||||
|
||||
Lemma choose_f_simpl: forall (D:cpo) e1 e2, choose D e1 e2 false = e2.
|
||||
intros D e1 e2. unfold choose. auto.
|
||||
Qed.
|
||||
|
||||
Lemma chooseVal: forall (D:cpo) b (e1 e2:DL D) v , choose _ e1 e2 b == Val v ->
|
||||
(b = true /\ e1 == Val v) \/ (b = false /\ e2 == Val v).
|
||||
intros D b e1 e2 v.
|
||||
assert (choose _ e1 e2 b = fcontit (choosec e1 e2) b). auto.
|
||||
rewrite H. simpl. clear H.
|
||||
destruct b. intros e1val. left. split ; auto.
|
||||
intros ; right ; split ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma SemEnvCons : forall t E (env : SemEnv (t :: E)), exists d, exists ds, env = PAIR ds d.
|
||||
intros. dependent inversion env. exists t1. exists t0. auto.
|
||||
Qed.
|
||||
|
||||
Lemma Discrete_injective : forall t (v1 v2:Discrete t), v1 == v2 -> v1 = v2.
|
||||
intros t v1 v2 eq.
|
||||
destruct eq. destruct H. destruct H0. trivial.
|
||||
Qed.
|
||||
|
||||
Lemma relEnv_ext env senv s s' : (forall x y, s x y = s' x y) -> relEnv env senv s -> relEnv env senv s'.
|
||||
intros E env s s' C r. induction E. auto.
|
||||
simpl. simpl in r. destruct r as [rl rr].
|
||||
split. rewrite SND_simpl. rewrite SND_simpl in rl. unfold hdMap. unfold hdMap in rl.
|
||||
simpl. simpl in rl. rewrite <- (C a (ZVAR E a)). apply rl.
|
||||
rewrite FST_simpl. unfold tlMap. simpl. rewrite FST_simpl in rr. unfold tlMap in rr. simpl in rr.
|
||||
refine (IHE _ _ _ _ rr). intros x y. specialize (C x (SVAR a y)). auto.
|
||||
Qed.
|
||||
|
||||
Theorem FundamentalTheorem:
|
||||
(forall env ty v senv s, relEnv env senv s -> relVal ty (SemVal v senv) (substVal s v))
|
||||
/\
|
||||
(forall env ty e senv s, relEnv env senv s -> liftRel (relVal ty) (SemExp e senv) (substExp s e)).
|
||||
Proof.
|
||||
apply ExpVal_ind.
|
||||
|
||||
(* TINT *)
|
||||
simpl; auto.
|
||||
|
||||
(* TBOOL *)
|
||||
simpl; auto.
|
||||
|
||||
(* TVAR *)
|
||||
intros E t v env s Renv.
|
||||
induction v.
|
||||
destruct (SemEnvCons env) as [d [ds eq]]. subst. simpl. fold SemEnv in *.
|
||||
simpl in Renv. destruct Renv as [R1 _]. trivial.
|
||||
destruct (SemEnvCons env) as [d [ds eq]]. subst. simpl. fold SemEnv in *.
|
||||
simpl in Renv. destruct Renv as [_ R2]. rewrite fcont_comp_simpl.
|
||||
specialize (IHv ds (tlMap s) R2). auto.
|
||||
|
||||
(* TFIX *)
|
||||
intros E t1 t2 e IH env s R.
|
||||
simpl SemVal.
|
||||
rewrite fcont_comp_simpl.
|
||||
rewrite substTFIX.
|
||||
rewrite FIXP_simpl. rewrite Fixp_simpl.
|
||||
apply (@fixp_ind _ _ (fcontit (curry (curry (SemExp e)) env))).
|
||||
refine (@rel_admissible (t1 --> t2) _).
|
||||
simpl. exists (substExp (liftSubst t1 (liftSubst (t1 --> t2) s)) e) ; split ; auto.
|
||||
intros d1 v1 rv1. unfold liftRel. intros dd. rewrite K_simpl.
|
||||
intros CC. assert False as F by (apply (PBot_val_incon CC)). inversion F.
|
||||
|
||||
intros f rf.
|
||||
exists (substExp (liftSubst t1 (liftSubst (t1 --> t2) s)) e). split ; auto.
|
||||
intros d1 v1 rv1. unfold liftRel. intros dd fd.
|
||||
assert (SemExp e (env,f,d1) == Val dd) as se by auto. clear fd.
|
||||
rewrite <- (proj2 substComposeSubst).
|
||||
rewrite <- substTFIX.
|
||||
assert (relEnv ((t1 :: t1 --> t2 :: E)) (env,f,d1) (consMap v1 (consMap (substVal (env':=nil) s (TFIX e)) s))).
|
||||
simpl relEnv. rewrite SND_simpl. simpl. split.
|
||||
unfold hdMap. simpl. apply rv1. split.
|
||||
unfold tlMap. unfold hdMap. simpl. rewrite substTFIX. exists (substExp
|
||||
(liftSubst t1 (liftSubst (t1 --> t2) s)) e). split ; auto.
|
||||
intros d2 v2 rv2. rewrite FST_simpl. rewrite SND_simpl. simpl.
|
||||
simpl in rf. destruct rf as [bb [bbeq Pb]].
|
||||
assert (bb = (substExp (liftSubst t1 (liftSubst (t1 --> t2) s)) e)) as beq.
|
||||
apply (TFIX_injective (sym_equal bbeq)).
|
||||
rewrite beq in *. clear bb beq bbeq. specialize (Pb d2 v2 rv2). apply Pb.
|
||||
|
||||
rewrite FST_simpl. rewrite FST_simpl. simpl. unfold tlMap. simpl.
|
||||
refine (relEnv_ext _ R). intros x y ; auto.
|
||||
|
||||
specialize (IH _ _ H). unfold liftRel in IH. specialize (IH _ se).
|
||||
rewrite composeCons. rewrite composeCons.
|
||||
rewrite composeSubstIdLeft. apply IH.
|
||||
|
||||
(* TP *)
|
||||
intros E t1 t2 v1 IH1 v2 IH2 env s R.
|
||||
simpl. specialize (IH1 env s R). specialize (IH2 env s R).
|
||||
exists (substVal s v1). exists (substVal s v2). split; auto.
|
||||
|
||||
(* TFST *)
|
||||
intros E t1 t2 v IH env s R.
|
||||
simpl. specialize (IH env s R).
|
||||
unfold liftRel. intros d1 eq.
|
||||
rewrite substTFST.
|
||||
destruct (ClosedPair (substVal s v)) as [v1 [v2 veq]].
|
||||
rewrite veq in *. clear veq.
|
||||
exists v1. split. apply e_Fst. simpl relVal in IH.
|
||||
destruct IH as [v3 [v4 [veq [gl1 gl2]]]]. destruct (TPAIR_injective veq) as [v1eq v2eq]. subst. clear veq.
|
||||
rewrite fcont_comp_simpl in eq.
|
||||
rewrite fcont_comp_simpl in eq.
|
||||
rewrite eta_val in eq.
|
||||
assert (H := vinj eq).
|
||||
apply (relVal_lower (proj2 H) gl1).
|
||||
|
||||
(* TSND *)
|
||||
intros E t1 t2 v IH env s R.
|
||||
simpl. fold SemTy. specialize (IH env s R).
|
||||
unfold liftRel. intros d2 eq. rewrite fcont_comp_simpl in eq. rewrite fcont_comp_simpl in eq.
|
||||
rewrite substTSND.
|
||||
destruct (ClosedPair (substVal s v)) as [v1 [v2 veq]].
|
||||
rewrite veq in *. clear veq.
|
||||
exists v2. split. apply e_Snd. simpl relVal in IH.
|
||||
destruct IH as [v3 [v4 [veq [gl1 gl2]]]]. destruct (TPAIR_injective veq) as [v1eq v2eq]. subst. clear veq.
|
||||
rewrite eta_val in eq.
|
||||
assert (H := vinj eq).
|
||||
apply (relVal_lower (proj2 H) gl2).
|
||||
|
||||
(* TOP *)
|
||||
intros E op v1 IH1 v2 IH2 env s R.
|
||||
simpl. specialize (IH1 env s R). specialize (IH2 env s R).
|
||||
unfold liftRel. intros d eq. rewrite fcont_comp_simpl in eq. rewrite fcont_comp_simpl in eq. rewrite eta_val in eq. assert (H := vinj eq). clear eq.
|
||||
rewrite substTOP.
|
||||
destruct (ClosedInt (substVal s v1)) as [i1 eq1].
|
||||
destruct (ClosedInt (substVal s v2)) as [i2 eq2].
|
||||
rewrite eq1 in *. rewrite eq2 in *. clear eq1 eq2.
|
||||
exists (TINT (op i1 i2)).
|
||||
split. apply e_Op. rewrite PROD_fun_simpl in H. rewrite uncurry_simpl in H.
|
||||
unfold relVal in *. inversion IH1. inversion IH2. clear IH1 IH2 H1 H2.
|
||||
assert ((op (SemVal v1 env) (SemVal v2 env) : Discrete nat) == d) by auto. clear H.
|
||||
unfold SemTy in d.
|
||||
assert (H := Discrete_injective H0). rewrite <- H. trivial.
|
||||
|
||||
(* TGT *)
|
||||
intros E v1 IH1 v2 IH2 env s R.
|
||||
simpl. specialize (IH1 env s R). specialize (IH2 env s R).
|
||||
unfold liftRel. intros d eq. rewrite fcont_comp_simpl in eq. rewrite fcont_comp_simpl in eq. rewrite eta_val in eq. assert (H := vinj eq). clear eq.
|
||||
rewrite substTGT.
|
||||
destruct (ClosedInt (substVal s v1)) as [i1 eq1].
|
||||
destruct (ClosedInt (substVal s v2)) as [i2 eq2].
|
||||
rewrite eq1 in *. rewrite eq2 in *. clear eq1 eq2.
|
||||
exists (TBOOL (ble_nat i2 i1)).
|
||||
split. apply e_Gt.
|
||||
rewrite PROD_fun_simpl in H. rewrite uncurry_simpl in H.
|
||||
unfold relVal in *. inversion IH1. inversion IH2. clear IH1 IH2 H1 H2.
|
||||
assert ((ble_nat (SemVal v2 env) (SemVal v1 env) : Discrete bool) == d) by auto. clear H.
|
||||
unfold SemTy in d.
|
||||
assert (H := Discrete_injective H0). rewrite <- H. trivial.
|
||||
|
||||
|
||||
(* VAL *)
|
||||
intros E t v IH env s R.
|
||||
simpl. specialize (IH env s R).
|
||||
unfold liftRel. intros d eq. rewrite fcont_comp_simpl in eq. rewrite eta_val in eq. assert (H := vinj eq). clear eq.
|
||||
exists (substVal s v).
|
||||
split. apply e_Val.
|
||||
apply (relVal_lower (proj2 H) IH).
|
||||
|
||||
(* TLET *)
|
||||
intros E t1 t2 e1 IH1 e2 IH2 env s R.
|
||||
unfold liftRel. intros d2 d2eq.
|
||||
specialize (IH1 env s R). unfold liftRel in *.
|
||||
simpl in d2eq.
|
||||
|
||||
rewrite substTLET. rewrite fcont_comp_simpl in d2eq. rewrite PROD_fun_simpl in d2eq.
|
||||
rewrite ID_simpl in d2eq. simpl in d2eq.
|
||||
destruct (KLEISLIR_ValVal d2eq) as [d1 [d1eq d2veq]]. clear d2eq.
|
||||
specialize (IH1 _ d1eq).
|
||||
destruct IH1 as [v1 [ev1 R1]].
|
||||
specialize (IH2 (env,d1) (consMap v1 s)).
|
||||
assert (RR : relEnv (t1::E) (env,d1) (consMap v1 s)).
|
||||
simpl relEnv. split.
|
||||
rewrite SND_simpl. assumption.
|
||||
rewrite tlConsMap.
|
||||
rewrite FST_simpl. assumption.
|
||||
specialize (IH2 RR d2 d2veq).
|
||||
destruct IH2 as [v2 [ev2 R2]].
|
||||
exists v2. split.
|
||||
refine (e_Let ev1 _).
|
||||
rewrite <- (proj2 substComposeSubst). rewrite composeCons. rewrite composeSubstIdLeft. apply ev2. assumption.
|
||||
|
||||
(* TAPP *)
|
||||
intros E t1 t2 v1 IH1 v2 IH2 env s R.
|
||||
specialize (IH1 env s R). specialize (IH2 env s R).
|
||||
unfold liftRel. intros d eq. simpl SemExp in eq. fold SemTy in eq. rewrite fcont_comp_simpl in eq.
|
||||
rewrite substTAPP.
|
||||
destruct (ClosedFunction (substVal s v1)) as [e1 eq1].
|
||||
unfold relVal in IH1. fold relVal in IH1. destruct IH1 as [e1' [eq' R']].
|
||||
rewrite eq' in eq1.
|
||||
assert (H := TFIX_injective eq1).
|
||||
subst. clear eq1.
|
||||
specialize (R' (SemVal v2 env) (substVal s v2) IH2).
|
||||
clear IH2.
|
||||
unfold liftRel in R'. specialize (R' d eq).
|
||||
destruct R' as [v [ev Rv]].
|
||||
exists v. rewrite eq' in ev.
|
||||
assert (H := e_App ev).
|
||||
split. simpl. rewrite eq'. apply H.
|
||||
apply Rv.
|
||||
|
||||
(* TIF *)
|
||||
intros E t v IH e1 IH1 e2 IH2 env s R.
|
||||
specialize (IH env s R). specialize (IH1 env s R). specialize (IH2 env s R).
|
||||
unfold liftRel in *. intros d eq. simpl SemExp in eq.
|
||||
rewrite substTIF.
|
||||
destruct (ClosedBool (substVal s v)) as [b beq].
|
||||
rewrite beq in IH.
|
||||
unfold relVal in IH.
|
||||
inversion IH. clear IH.
|
||||
simpl substExp. rewrite beq.
|
||||
rewrite fcont_comp3_simpl in eq.
|
||||
case (chooseVal eq).
|
||||
(* chooseVal eq = true *)
|
||||
intros [H1 H2].
|
||||
destruct b.
|
||||
(* b=true *)
|
||||
specialize (IH1 d H2).
|
||||
destruct IH1 as [v' [ev Rv]].
|
||||
exists v'.
|
||||
split.
|
||||
apply (e_IfTrue (substExp s e2) ev).
|
||||
apply Rv.
|
||||
(* b=false *)
|
||||
rewrite H1 in H0. inversion H0.
|
||||
(* chooseVal eq = false *)
|
||||
intros [H1 H2].
|
||||
destruct b.
|
||||
(* b=true *)
|
||||
rewrite H1 in H0. inversion H0.
|
||||
(* b=false *)
|
||||
specialize (IH2 d H2).
|
||||
destruct IH2 as [v' [ev Rv]].
|
||||
exists v'.
|
||||
split.
|
||||
apply (e_IfFalse (substExp s e1) ev).
|
||||
apply Rv.
|
||||
Qed.
|
||||
|
||||
Corollary Adequacy: forall ty (e : CExp ty) d, SemExp e tt == Val d -> exists v, e =>> v.
|
||||
Proof.
|
||||
intros t e d val.
|
||||
destruct FundamentalTheorem as [_ FT].
|
||||
specialize (FT nil t e tt (idSubst nil)).
|
||||
assert (relEnv nil (tt:DOne) (idSubst nil)). unfold relEnv. auto. specialize (FT H).
|
||||
unfold liftRel in FT. specialize (FT d val).
|
||||
destruct FT as [v [ev _]]. exists v. rewrite (proj2 applyIdSubst) in ev. trivial.
|
||||
Qed.
|
||||
|
255
papers/domains-2009/typeddensem.v
Executable file
255
papers/domains-2009/typeddensem.v
Executable file
|
@ -0,0 +1,255 @@
|
|||
(* typeddensem.v
|
||||
denotational semantics
|
||||
This is the second attempt, this time using predomains with lifing
|
||||
*)
|
||||
|
||||
Require Import utility.
|
||||
Require Export typedlambda.
|
||||
Require Import PredomAll.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
Open Scope O_scope.
|
||||
|
||||
(*==========================================================================
|
||||
Meaning of types and contexts
|
||||
==========================================================================*)
|
||||
|
||||
Delimit Scope Ty_scope with Ty.
|
||||
|
||||
Open Scope Ty_scope.
|
||||
|
||||
Fixpoint SemTy ty :=
|
||||
match ty with
|
||||
| Int => Discrete nat
|
||||
| Bool => Discrete bool
|
||||
| ty1 --> ty2 => SemTy ty1 -C-> (SemTy ty2) _BOT
|
||||
| ty1 ** ty2 => SemTy ty1 *c* SemTy ty2
|
||||
end.
|
||||
|
||||
(* We define the semantics of an environment "from the right";
|
||||
this avoids need for SWAP in the semantics *)
|
||||
Fixpoint SemEnv env :=
|
||||
match env with
|
||||
| nil => DOne
|
||||
| ty :: env => SemEnv env *c* SemTy ty
|
||||
end.
|
||||
|
||||
(* Lookup in an environment *)
|
||||
Fixpoint SemVar env ty (var : Var env ty) : SemEnv env -c> SemTy ty :=
|
||||
match var with
|
||||
| ZVAR _ _ => SND
|
||||
| SVAR _ _ _ v => SemVar v << FST
|
||||
end.
|
||||
|
||||
(* A coercion to allow embedding values of type X into its lifted
|
||||
discrete domain. DL_discrete is funny because we need to put in the
|
||||
stuff the existing coercions are already adding or else the matching
|
||||
will fail. The Id type operator is there because coercions are by
|
||||
name. *)
|
||||
|
||||
Definition Id (x : Type) := x.
|
||||
Definition DL_discrete X := tord (tcpo (DL (Discrete X))).
|
||||
|
||||
Definition discrete_lift : forall X : Set, Id X -> DL_discrete X :=
|
||||
fun X x => eta (x : Discrete X).
|
||||
|
||||
Implicit Arguments discrete_lift [X].
|
||||
|
||||
Coercion discrete_lift : Id >-> DL_discrete.
|
||||
|
||||
(* The need to use the defined operator names -- especially Id --
|
||||
makes it seem like a marginal benefit. But maybe it's worth
|
||||
it as an example of how coercions work.
|
||||
|
||||
Definition top : DL_discrete unit := (tt : Id unit).
|
||||
Definition top': DL (Discrete unit) := discrete_lift tt.
|
||||
|
||||
*)
|
||||
Definition indiscrete (X:Type) (x:X) : Discrete X := x.
|
||||
|
||||
Definition SimpleOpm (A:Type) (B:cpo) (op : A -> ((tord B) : Type)) : Discrete A -m> B.
|
||||
intros.
|
||||
exists op.
|
||||
unfold monotonic.
|
||||
intros.
|
||||
simpl in *.
|
||||
rewrite H. auto.
|
||||
Defined.
|
||||
|
||||
Definition SimpleOp (A:Set) (B:cpo) (op : A -> B) : Discrete A -c> B.
|
||||
intros.
|
||||
exists (SimpleOpm op).
|
||||
unfold continuous.
|
||||
intros c.
|
||||
assert (c 0 = lub c). auto.
|
||||
rewrite <- H. simpl.
|
||||
assert (op (c 0) <= (SimpleOpm op @ c) 0). auto.
|
||||
eapply Ole_trans. apply H0.
|
||||
apply le_lub.
|
||||
Defined.
|
||||
|
||||
Definition SimpleOp2mm (A B C : Type) (op : A -> B -> C) : Discrete A -> Discrete B -m> Discrete C.
|
||||
intros.
|
||||
exists (op X).
|
||||
unfold monotonic.
|
||||
intros.
|
||||
simpl in *.
|
||||
rewrite H. auto.
|
||||
Defined.
|
||||
|
||||
Definition SimpleOp2c A B C (op:A -> B -> C) :
|
||||
Discrete A -> Discrete B -c> Discrete C.
|
||||
intros.
|
||||
exists (SimpleOp2mm op X).
|
||||
simpl in *.
|
||||
unfold continuous.
|
||||
intros.
|
||||
simpl. auto.
|
||||
Defined.
|
||||
|
||||
Definition SimpleOp2m (A B C:Type) (op : A -> B -> C) :
|
||||
Discrete A -m> Discrete B -C-> Discrete C.
|
||||
intros.
|
||||
exists (SimpleOp2c op).
|
||||
unfold monotonic.
|
||||
intros.
|
||||
simpl in *.
|
||||
rewrite H. auto.
|
||||
Defined.
|
||||
|
||||
Definition SimpleOp2 A B C (op:A -> B -> C) :
|
||||
Discrete A -c> Discrete B -C-> Discrete C.
|
||||
intros.
|
||||
exists (SimpleOp2m op).
|
||||
unfold continuous ; intros ; simpl ; auto.
|
||||
Defined.
|
||||
|
||||
Definition choosemm : forall (D : cpo), D -> D -> Discrete bool -m> D.
|
||||
intros D X Y.
|
||||
exists (fun (b:bool) => if b then X else Y).
|
||||
unfold monotonic ; intros ; simpl in *; rewrite H ; auto.
|
||||
Defined.
|
||||
|
||||
Definition choosec : forall (D : cpo), D -> D -> Discrete bool -c> D.
|
||||
intros D X Y.
|
||||
exists (choosemm X Y).
|
||||
unfold continuous.
|
||||
intros c ; simpl in *.
|
||||
assert ((if c 0 then X else Y) = ((choosemm X Y @ c) 0)).
|
||||
simpl ; auto. rewrite H. apply le_lub.
|
||||
Defined.
|
||||
|
||||
Definition choosem : forall (D : cpo), D -> D -m> Discrete bool -C-> D.
|
||||
intros.
|
||||
exists (@choosec D X).
|
||||
unfold monotonic.
|
||||
intros.
|
||||
simpl. intro.
|
||||
destruct x0 ; auto.
|
||||
Defined.
|
||||
|
||||
Definition choosecc : forall (D : cpo), D -> D -c> Discrete bool -C-> D.
|
||||
intros.
|
||||
exists (@choosem D X).
|
||||
unfold continuous.
|
||||
intros c.
|
||||
simpl.
|
||||
intro.
|
||||
destruct x.
|
||||
eapply Ole_trans.
|
||||
assert (X <= ((Fcontit (Discrete bool) D @ (choosem X @ c) <_> true) 0)).
|
||||
simpl ; auto. apply H.
|
||||
apply le_lub.
|
||||
refine (lub_le_compat _).
|
||||
refine (fmon_le_intro _).
|
||||
intros n.
|
||||
auto.
|
||||
Defined.
|
||||
|
||||
Definition choosecm : forall (D:cpo), D -m> D -C-> Discrete bool -C-> D.
|
||||
intros.
|
||||
exists (@choosecc D).
|
||||
unfold monotonic.
|
||||
intros.
|
||||
simpl.
|
||||
intros.
|
||||
destruct x1. auto. auto.
|
||||
Defined.
|
||||
|
||||
Definition choose : forall (D:cpo), D -c> D -C-> Discrete bool -C-> D.
|
||||
intros.
|
||||
exists (@choosecm D).
|
||||
unfold continuous.
|
||||
intros c.
|
||||
simpl. intros.
|
||||
destruct x0.
|
||||
refine (lub_le_compat (fmon_le_intro _)).
|
||||
intros n. auto.
|
||||
|
||||
eapply Ole_trans.
|
||||
assert (x <= ((Fcontit (Discrete bool) D @ ((Fcontit D (Discrete bool -C-> D) @ (choosecm D @ c)) <_> x)) <_> false) 0).
|
||||
simpl ; auto. apply H.
|
||||
apply le_lub.
|
||||
Defined.
|
||||
|
||||
Canonical Structure Discrete.
|
||||
|
||||
(*==========================================================================
|
||||
Meaning of values and expressions
|
||||
==========================================================================*)
|
||||
|
||||
Fixpoint SemExp env ty (e : Exp env ty) : SemEnv env -c> (SemTy ty) _BOT :=
|
||||
match e with
|
||||
| TOP _ op v1 v2 => eta << uncurry (SimpleOp2 op) << <| SemVal v1 , SemVal v2 |>
|
||||
| TGT _ v1 v2 => eta << uncurry (SimpleOp2 ble_nat) << <| SemVal v2 , SemVal v1 |>
|
||||
| TAPP _ _ _ v1 v2 => EV << <| SemVal v1 , SemVal v2 |>
|
||||
| TVAL _ _ v => eta << SemVal v
|
||||
| TLET _ _ _ e1 e2 => KLEISLIR (SemExp e2) << <| ID , SemExp e1 |>
|
||||
| TIF _ _ v e1 e2 => (choose _ @3_ (SemExp e1)) (SemExp e2) (SemVal v)
|
||||
| TFST _ _ _ v => eta << FST << SemVal v
|
||||
| TSND _ _ _ v => eta << SND << SemVal v
|
||||
end
|
||||
|
||||
with SemVal env ty (v : Value env ty) : SemEnv env -c> SemTy ty :=
|
||||
match v with
|
||||
| TINT _ n => K _ (n : Discrete nat)
|
||||
| TBOOL _ b => K _ (b : Discrete bool)
|
||||
| TVAR _ _ i => SemVar i
|
||||
| TFIX _ _ _ e => FIXP << curry (curry (SemExp e))
|
||||
| TPAIR _ _ _ v1 v2 => <| SemVal v1 , SemVal v2 |>
|
||||
end.
|
||||
|
||||
|
||||
Lemma choosec_t_simpl: forall (D:cpo) (a b:D), choosec a b true = a.
|
||||
intros D a b.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma choosec_f_simpl: forall (D:cpo) (a b:D), choosec a b false = b.
|
||||
intros D a b.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma choose_t_simpl: forall D a b, choose D a b true = a.
|
||||
auto.
|
||||
Qed.
|
||||
Lemma choose_f_simpl: forall D a b, choose D a b false = b.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism E t (e:Exp E t) : (SemExp e)
|
||||
with signature (@Oeq (SemEnv E)) ==> (@Oeq (DL (SemTy t)))
|
||||
as SemExp_eq_compat.
|
||||
intros e0 e1 eeq.
|
||||
destruct eeq. split ; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism E t (e:Value E t) : (SemVal e)
|
||||
with signature (@Oeq (SemEnv E)) ==> (@Oeq (SemTy t))
|
||||
as SemVal_eq_compat.
|
||||
intros e0 e1 eeq.
|
||||
destruct eeq. split ; auto.
|
||||
Qed.
|
||||
|
509
papers/domains-2009/typedlambda.v
Executable file
509
papers/domains-2009/typedlambda.v
Executable file
|
@ -0,0 +1,509 @@
|
|||
(*==========================================================================
|
||||
Call-by-value simply-typed lambda-calculus with recursion, pairs, and arithmetic
|
||||
==========================================================================*)
|
||||
|
||||
Require Import utility.
|
||||
Require Import List.
|
||||
Require Import Program.Equality.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Set Printing Implicit Defensive.
|
||||
|
||||
(*====== coqdoc printing directives ========================================*)
|
||||
(** printing --> %\ensuremath{\mathrel{\texttt{->}}}% *)
|
||||
(** printing ** %\ensuremath{\mathrel{\texttt{*}}}% *)
|
||||
(** printing Int %{\textsf{Int}}% *)
|
||||
(** printing Bool %{\textsf{Bool}}% *)
|
||||
(** printing Arrow %{\textsf{Arrow}}% *)
|
||||
(** printing Prod %{\textsf{Prod}}% *)
|
||||
(** printing Ty %{\textsf{Ty}}% *)
|
||||
(** printing Env %{\textsf{Env}}% *)
|
||||
|
||||
(** printing env %{\ensuremath{\Gamma}}% *)
|
||||
(** printing env' %{\ensuremath{\Gamma'}}% *)
|
||||
(** printing ty %{\ensuremath{\tau}}% *)
|
||||
(** printing ty1 %\ensuremath{\tau_1}% *)
|
||||
(** printing ty2 %\ensuremath{\tau_2}% *)
|
||||
(** printing ty' %\ensuremath{\tau'}% *)
|
||||
|
||||
(*==========================================================================
|
||||
Types and contexts
|
||||
==========================================================================*)
|
||||
|
||||
Inductive Ty := Int | Bool | Arrow (ty1 ty2 : Ty) | Prod (ty1 ty2 : Ty).
|
||||
|
||||
Infix " --> " := Arrow. (* (at level 55). *)
|
||||
Infix " ** " := Prod (at level 55).
|
||||
|
||||
Definition Env := list Ty.
|
||||
|
||||
(* begin hide *)
|
||||
Notation "[ x , .. , y ]" := (cons x .. (cons y nil) ..) : list_scope.
|
||||
(* end hide *)
|
||||
|
||||
(*==========================================================================
|
||||
Typed terms in context
|
||||
==========================================================================*)
|
||||
|
||||
Inductive Var : Env -> Ty -> Type :=
|
||||
| ZVAR : forall env ty, Var (ty :: env) ty
|
||||
| SVAR : forall env ty ty', Var env ty -> Var (ty' :: env) ty.
|
||||
|
||||
Inductive Value : Env -> Ty -> Type :=
|
||||
| TINT : forall env, nat -> Value env Int
|
||||
| TBOOL : forall env, bool -> Value env Bool
|
||||
| TVAR :> forall env ty, Var env ty -> Value env ty
|
||||
| TFIX : forall env ty1 ty2, Exp (ty1 :: ty1 --> ty2 :: env) ty2 -> Value env (ty1 --> ty2)
|
||||
| TPAIR : forall env ty1 ty2, Value env ty1 -> Value env ty2 -> Value env (ty1 ** ty2)
|
||||
with Exp : Env -> Ty -> Type :=
|
||||
| TFST : forall env ty1 ty2, Value env (ty1 ** ty2) -> Exp env ty1
|
||||
| TSND : forall env ty1 ty2, Value env (ty1 ** ty2) -> Exp env ty2
|
||||
| TOP : forall env, (nat -> nat -> nat) -> Value env Int -> Value env Int -> Exp env Int
|
||||
| TGT : forall env, Value env Int -> Value env Int -> Exp env Bool
|
||||
| TVAL : forall env ty, Value env ty -> Exp env ty
|
||||
| TLET : forall env ty1 ty2, Exp env ty1 -> Exp (ty1 :: env) ty2 -> Exp env ty2
|
||||
| TAPP : forall env ty1 ty2, Value env (ty1 --> ty2) -> Value env ty1 -> Exp env ty2
|
||||
| TIF : forall env ty, Value env Bool -> Exp env ty -> Exp env ty -> Exp env ty.
|
||||
|
||||
Implicit Arguments TBOOL [env].
|
||||
Implicit Arguments TINT [env].
|
||||
|
||||
(* begin hide *)
|
||||
Scheme Val_rec2 := Induction for Value Sort Set
|
||||
with Exp_rec2 := Induction for Exp Sort Set.
|
||||
|
||||
Scheme Val_type2 := Induction for Value Sort Type
|
||||
with Exp_type2 := Induction for Exp Sort Type.
|
||||
|
||||
Scheme Val_ind2 := Induction for Value Sort Prop
|
||||
with Exp_ind2 := Induction for Exp Sort Prop.
|
||||
|
||||
Combined Scheme ExpVal_ind from Val_ind2, Exp_ind2.
|
||||
(* end hide *)
|
||||
|
||||
(* Closed expressions and values *)
|
||||
Definition CExp ty := Exp nil ty.
|
||||
Definition CValue ty := Value nil ty.
|
||||
|
||||
(*==========================================================================
|
||||
Variable-domain maps.
|
||||
By instantiating P with Var we get renamings.
|
||||
By instantiating P with Val we get substitutions.
|
||||
==========================================================================*)
|
||||
|
||||
Definition Map (P:Env -> Ty -> Type) E E' := forall t, Var E t -> P E' t.
|
||||
|
||||
(* Head, tail and cons *)
|
||||
Definition tlMap P E E' t (m:Map P (t::E) E') : Map P E E' := fun t' v => m t' (SVAR t v).
|
||||
Definition hdMap P E E' t (m:Map P (t::E) E') : P E' t := m t (ZVAR _ _).
|
||||
Implicit Arguments tlMap [P E E' t].
|
||||
Implicit Arguments hdMap [P E E' t].
|
||||
|
||||
Program Definition consMap P E E' t (v:P E' t) (m:Map P E E') : Map P (t::E) E' :=
|
||||
fun t' (var:Var (t::E) t') =>
|
||||
match var with
|
||||
| ZVAR _ _ => v
|
||||
| SVAR _ _ _ var => m _ var
|
||||
end.
|
||||
Implicit Arguments consMap [P E E' t].
|
||||
|
||||
Axiom MapExtensional : forall P E E' (r1 r2 : Map P E E'), (forall t var, r1 t var = r2 t var) -> r1 = r2.
|
||||
|
||||
Lemma hdConsMap : forall P env env' ty (v : P env' ty) (s : Map P env env'), hdMap (consMap v s) = v. Proof. auto. Qed.
|
||||
Lemma tlConsMap : forall P env env' ty (v : P env' ty) (s : Map P env env'), tlMap (consMap v s) = s. Proof. intros. apply MapExtensional. auto. Qed.
|
||||
|
||||
Lemma consMapInv : forall P env env' ty (m:Map P (ty :: env) env'), exists m', exists v, m = consMap v m'.
|
||||
intros. exists (tlMap m). exists (hdMap m).
|
||||
apply MapExtensional. dependent destruction var; auto.
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Package of operations used with a Map
|
||||
vr maps a Var into Var or Value (so is either the identity or TVAR)
|
||||
vl maps a Var or Value to a Value (so is either TVAR or the identity)
|
||||
wk weakens a Var or Value (so is either SVAR or renaming through SVAR on a value)
|
||||
==========================================================================*)
|
||||
Record MapOps (P : Env -> Ty -> Type) :=
|
||||
{
|
||||
vr : forall env ty, Var env ty -> P env ty;
|
||||
vl : forall env ty, P env ty -> Value env ty;
|
||||
wk : forall env ty ty', P env ty -> P (ty' :: env) ty
|
||||
}.
|
||||
|
||||
Section MapOps.
|
||||
|
||||
Variable P : Env -> Ty -> Type.
|
||||
Variable ops : MapOps P.
|
||||
|
||||
Program Definition lift env env' ty (m : Map P env env') : Map P (ty :: env) (ty :: env') :=
|
||||
fun ty' => fun var => match var with
|
||||
| ZVAR _ _ => vr ops (ZVAR _ _)
|
||||
| SVAR _ _ _ x => wk ops _ (m _ x)
|
||||
end.
|
||||
Implicit Arguments lift [env env'].
|
||||
|
||||
Definition shiftMap env env' ty (m : Map P env env') : Map P env (ty :: env') := fun ty' => fun var => wk ops _ (m ty' var).
|
||||
Implicit Arguments shiftMap [env env'].
|
||||
|
||||
Lemma shiftConsMap : forall env env' ty (m : Map P env env') (x : P env' ty) ty', shiftMap ty' (consMap x m) = consMap (wk ops _ x) (shiftMap ty' m).
|
||||
Proof.
|
||||
intros env env' ty m x ty'. apply MapExtensional.
|
||||
intros ty0 var0. dependent destruction var0; auto.
|
||||
Qed.
|
||||
|
||||
Fixpoint travVal env env' ty (v : Value env ty) : Map P env env' -> Value env' ty :=
|
||||
match v with
|
||||
| TINT _ i => fun m => TINT i
|
||||
| TBOOL _ b => fun m => TBOOL b
|
||||
| TVAR _ _ v => fun m => vl ops (m _ v)
|
||||
| TFIX _ _ _ e => fun m => TFIX (travExp e (lift _ (lift _ m)))
|
||||
| TPAIR _ _ _ e1 e2 => fun m => TPAIR (travVal e1 m) (travVal e2 m)
|
||||
end
|
||||
with travExp env env' ty (e : Exp env ty) : Map P env env' -> Exp env' ty :=
|
||||
match e with
|
||||
| TOP _ op e1 e2 => fun m => TOP op (travVal e1 m) (travVal e2 m)
|
||||
| TGT _ e1 e2 => fun m => TGT (travVal e1 m) (travVal e2 m)
|
||||
| TFST _ _ _ e => fun m => TFST (travVal e m)
|
||||
| TSND _ _ _ e => fun m => TSND (travVal e m)
|
||||
| TVAL _ _ v => fun m => TVAL (travVal v m)
|
||||
| TIF _ _ v e1 e2 => fun m => TIF (travVal v m) (travExp e1 m) (travExp e2 m)
|
||||
| TAPP _ _ _ e1 e2 => fun m => TAPP (travVal e1 m) (travVal e2 m)
|
||||
| TLET _ _ _ e1 e2 => fun m => TLET (travExp e1 m) (travExp e2 (lift _ m))
|
||||
end.
|
||||
|
||||
Definition mapVal env env' ty m v := @travVal env env' ty v m.
|
||||
Definition mapExp env env' ty m e := @travExp env env' ty e m.
|
||||
|
||||
Variable env env' : Env.
|
||||
Variable m : Map P env env'.
|
||||
Variable ty1 ty2 : Ty.
|
||||
|
||||
Lemma mapTVAR : forall (var : Var _ ty1), mapVal m (TVAR var) = vl ops (m var). auto. Qed.
|
||||
Lemma mapTINT : forall n, mapVal m (TINT n) = TINT n. auto. Qed.
|
||||
Lemma mapTBOOL : forall n, mapVal m (TBOOL n) = TBOOL n. auto. Qed.
|
||||
Lemma mapTPAIR : forall (v1 : Value _ ty1) (v2 : Value _ ty2), mapVal m (TPAIR v1 v2) = TPAIR (mapVal m v1) (mapVal m v2). auto. Qed.
|
||||
Lemma mapTFST : forall (v : Value _ (ty1 ** ty2)), mapExp m (TFST v) = TFST (mapVal m v). auto. Qed.
|
||||
Lemma mapTSND : forall (v : Value _ (ty1 ** ty2)), mapExp m (TSND v) = TSND (mapVal m v). auto. Qed.
|
||||
Lemma mapTFIX : forall (e : Exp (ty1 :: ty1-->ty2 :: _) ty2), mapVal m (TFIX e) = TFIX (mapExp (lift _ (lift _ m)) e). auto. Qed.
|
||||
Lemma mapTOP : forall op v1 v2, mapExp m (TOP op v1 v2) = TOP op (mapVal m v1) (mapVal m v2). auto. Qed.
|
||||
Lemma mapTGT : forall v1 v2, mapExp m (TGT v1 v2) = TGT (mapVal m v1) (mapVal m v2). auto. Qed.
|
||||
Lemma mapTVAL : forall (v : Value _ ty1), mapExp m (TVAL v) = TVAL (mapVal m v). auto. Qed.
|
||||
Lemma mapTLET : forall (e1 : Exp _ ty1) (e2 : Exp _ ty2), mapExp m (TLET e1 e2) = TLET (mapExp m e1) (mapExp (lift _ m) e2). auto. Qed.
|
||||
Lemma mapTIF : forall v (e1 e2 : Exp _ ty1), mapExp m (TIF v e1 e2) = TIF (mapVal m v) (mapExp m e1) (mapExp m e2). auto. Qed.
|
||||
Lemma mapTAPP : forall (v1 : Value _ (ty1 --> ty2)) v2, mapExp m (TAPP v1 v2) = TAPP (mapVal m v1) (mapVal m v2). auto. Qed.
|
||||
|
||||
End MapOps.
|
||||
|
||||
Hint Rewrite mapTVAR mapTINT mapTBOOL mapTPAIR mapTFST mapTSND mapTFIX mapTOP mapTGT mapTVAL mapTLET mapTIF mapTAPP : mapHints.
|
||||
|
||||
Implicit Arguments lift [P env env'].
|
||||
Implicit Arguments shiftMap [P env env'].
|
||||
|
||||
(*==========================================================================
|
||||
Variable renamings: Map Var
|
||||
==========================================================================*)
|
||||
|
||||
Definition Renaming := Map Var.
|
||||
Definition RenamingMapOps := (Build_MapOps (fun _ _ v => v) TVAR SVAR).
|
||||
|
||||
Definition renameVal := mapVal RenamingMapOps.
|
||||
Definition renameExp := mapExp RenamingMapOps.
|
||||
Definition liftRenaming := lift RenamingMapOps.
|
||||
Implicit Arguments liftRenaming [env env'].
|
||||
Definition shiftRenaming := shiftMap RenamingMapOps.
|
||||
Implicit Arguments shiftRenaming [env env'].
|
||||
|
||||
Section RenamingDefs.
|
||||
|
||||
Variable env env' : Env.
|
||||
Variable r : Renaming env env'.
|
||||
Variable ty1 ty2 : Ty.
|
||||
|
||||
Lemma renameTVAR : forall (var : Var env ty1), renameVal r (TVAR var) = TVAR (r var). auto. Qed.
|
||||
Lemma renameTINT : forall n, renameVal r (TINT n) = TINT n. auto. Qed.
|
||||
Lemma renameTBOOL : forall n, renameVal r (TBOOL n) = TBOOL n. auto. Qed.
|
||||
Lemma renameTPAIR : forall (v1 : Value _ ty1) (v2 : Value _ ty2), renameVal r (TPAIR v1 v2) = TPAIR (renameVal r v1) (renameVal r v2). auto. Qed.
|
||||
Lemma renameTFST : forall (v : Value _ (ty1 ** ty2)), renameExp r (TFST v) = TFST (renameVal r v). auto. Qed.
|
||||
Lemma renameTSND : forall (v : Value _ (ty1 ** ty2)), renameExp r (TSND v) = TSND (renameVal r v). auto. Qed.
|
||||
Lemma renameTFIX : forall (e : Exp (ty1 :: ty1-->ty2 :: _) ty2), renameVal r (TFIX e) = TFIX (renameExp (liftRenaming _ (liftRenaming _ r)) e). auto. Qed.
|
||||
Lemma renameTOP : forall op v1 v2, renameExp r (TOP op v1 v2) = TOP op (renameVal r v1) (renameVal r v2). auto. Qed.
|
||||
Lemma renameTGT : forall v1 v2, renameExp r (TGT v1 v2) = TGT (renameVal r v1) (renameVal r v2). auto. Qed.
|
||||
Lemma renameTVAL : forall (v : Value env ty1), renameExp r (TVAL v) = TVAL (renameVal r v). auto. Qed.
|
||||
Lemma renameTLET : forall (e1 : Exp _ ty1) (e2 : Exp _ ty2), renameExp r (TLET e1 e2) = TLET (renameExp r e1) (renameExp (liftRenaming _ r) e2). auto. Qed.
|
||||
Lemma renameTIF : forall v (e1 e2 : Exp _ ty1), renameExp r (TIF v e1 e2) = TIF (renameVal r v) (renameExp r e1) (renameExp r e2). auto. Qed.
|
||||
Lemma renameTAPP : forall (v1 : Value _ (ty1-->ty2)) v2, renameExp r (TAPP v1 v2) = TAPP (renameVal r v1) (renameVal r v2). auto. Qed.
|
||||
|
||||
End RenamingDefs.
|
||||
|
||||
Hint Rewrite renameTVAR renameTINT renameTBOOL renameTPAIR renameTFST renameTSND renameTFIX renameTOP renameTGT renameTVAL renameTLET renameTIF renameTAPP : renameHints.
|
||||
|
||||
Lemma LiftRenamingDef : forall env env' (r : Renaming env' env) ty, liftRenaming _ r = consMap (ZVAR _ ty) (shiftRenaming _ r).
|
||||
intros. apply MapExtensional. auto. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Identity renaming
|
||||
==========================================================================*)
|
||||
|
||||
Definition idRenaming env : Renaming env env := fun ty (var : Var env ty) => var.
|
||||
Implicit Arguments idRenaming [].
|
||||
|
||||
Lemma liftIdRenaming : forall E t, liftRenaming _ (idRenaming E) = idRenaming (t::E).
|
||||
intros. apply MapExtensional.
|
||||
dependent destruction var; auto.
|
||||
Qed.
|
||||
|
||||
Lemma applyIdRenaming :
|
||||
(forall env ty (v : Value env ty), renameVal (idRenaming env) v = v)
|
||||
/\ (forall env ty (e : Exp env ty), renameExp (idRenaming env) e = e).
|
||||
Proof.
|
||||
apply ExpVal_ind;
|
||||
(intros; try (autorewrite with renameHints using try rewrite liftIdRenaming; try rewrite liftIdRenaming; try rewrite H; try rewrite H0; try rewrite H1); auto).
|
||||
Qed.
|
||||
|
||||
Lemma idRenamingDef : forall ty env, idRenaming (ty :: env) = consMap (ZVAR _ _) (shiftRenaming ty (idRenaming env)).
|
||||
intros. apply MapExtensional. intros. dependent destruction var; auto. Qed.
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Composition of renaming
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeRenaming P env env' env'' (m : Map P env' env'') (r : Renaming env env') : Map P env env'' := fun t var => m _ (r _ var).
|
||||
|
||||
Lemma liftComposeRenaming : forall P ops E env' env'' t (m:Map P env' env'') (r:Renaming E env'), lift ops t (composeRenaming m r) = composeRenaming (lift ops t m) (liftRenaming t r).
|
||||
intros. apply MapExtensional. intros t0 var.
|
||||
dependent destruction var; auto.
|
||||
Qed.
|
||||
|
||||
Lemma applyComposeRenaming :
|
||||
(forall env ty (v : Value env ty) P ops env' env'' (m:Map P env' env'') (s : Renaming env env'),
|
||||
mapVal ops (composeRenaming m s) v = mapVal ops m (renameVal s v))
|
||||
/\ (forall env ty (e : Exp env ty) P ops env' env'' (m:Map P env' env'') (s : Renaming env env'),
|
||||
mapExp ops (composeRenaming m s) e = mapExp ops m (renameExp s e)).
|
||||
Proof.
|
||||
apply ExpVal_ind; intros;
|
||||
autorewrite with mapHints renameHints; try rewrite liftComposeRenaming; try rewrite liftComposeRenaming; try rewrite <- H;
|
||||
try rewrite <- H0; try rewrite <- H1; auto.
|
||||
Qed.
|
||||
|
||||
Lemma composeRenamingAssoc :
|
||||
forall P env env' env'' env''' (m : Map P env'' env''') r' (r : Renaming env env'),
|
||||
composeRenaming m (composeRenaming r' r) = composeRenaming (composeRenaming m r') r.
|
||||
Proof. auto. Qed.
|
||||
|
||||
Lemma composeRenamingIdLeft : forall env env' (r : Renaming env env'), composeRenaming (idRenaming _) r = r.
|
||||
Proof. intros. apply MapExtensional. auto. Qed.
|
||||
|
||||
Lemma composeRenamingIdRight : forall P env env' (m:Map P env env'), composeRenaming m (idRenaming _) = m.
|
||||
Proof. intros. apply MapExtensional. auto. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Substitution
|
||||
==========================================================================*)
|
||||
|
||||
Definition Subst := Map Value.
|
||||
Definition SubstMapOps := Build_MapOps TVAR (fun _ _ v => v) (fun env ty ty' => renameVal (fun _ => SVAR ty')).
|
||||
|
||||
(* Convert a renaming into a substitution *)
|
||||
Definition renamingToSubst env env' (r : Renaming env env') : Subst env env' := fun ty => fun v => TVAR (r ty v).
|
||||
|
||||
Definition substVal := mapVal SubstMapOps.
|
||||
Definition substExp := mapExp SubstMapOps.
|
||||
|
||||
Definition shiftSubst := shiftMap SubstMapOps.
|
||||
Implicit Arguments shiftSubst [env env'].
|
||||
|
||||
Definition liftSubst := lift SubstMapOps.
|
||||
Implicit Arguments liftSubst [env env'].
|
||||
|
||||
Section SubstDefs.
|
||||
|
||||
Variable env env' : Env.
|
||||
Variable s : Subst env env'.
|
||||
Variable ty1 ty2 : Ty.
|
||||
|
||||
Lemma substTVAR : forall (var : Var env ty1), substVal s (TVAR var) = s var. auto. Qed.
|
||||
Lemma substTINT : forall n, substVal s (TINT n) = TINT n. auto. Qed.
|
||||
Lemma substTBOOL : forall n, substVal s (TBOOL n) = TBOOL n. auto. Qed.
|
||||
Lemma substTPAIR : forall (v1 : Value _ ty1) (v2:Value _ ty2), substVal s (TPAIR v1 v2) = TPAIR (substVal s v1) (substVal s v2). auto. Qed.
|
||||
Lemma substTFST : forall (v : Value _ (ty1 ** ty2)), substExp s (TFST v) = TFST (substVal s v). auto. Qed.
|
||||
Lemma substTSND : forall (v : Value _ (ty1 ** ty2)), substExp s (TSND v) = TSND (substVal s v). auto. Qed.
|
||||
Lemma substTFIX : forall (e : Exp (ty1 :: ty1-->ty2 :: _) ty2), substVal s (TFIX e) = TFIX (substExp (liftSubst _ (liftSubst _ s)) e). auto. Qed.
|
||||
Lemma substTOP : forall op v1 v2, substExp s (TOP op v1 v2) = TOP op (substVal s v1) (substVal s v2). auto. Qed.
|
||||
Lemma substTGT : forall v1 v2, substExp s (TGT v1 v2) = TGT (substVal s v1) (substVal s v2). auto. Qed.
|
||||
Lemma substTVAL : forall (v : Value _ ty1), substExp s (TVAL v) = TVAL (substVal s v). auto. Qed.
|
||||
Lemma substTLET : forall (e1 : Exp _ ty1) (e2 : Exp _ ty2), substExp s (TLET e1 e2) = TLET (substExp s e1) (substExp (liftSubst _ s) e2). auto. Qed.
|
||||
Lemma substTIF : forall v (e1 e2 : Exp _ ty1), substExp s (TIF v e1 e2) = TIF (substVal s v) (substExp s e1) (substExp s e2). auto. Qed.
|
||||
Lemma substTAPP : forall (v1:Value _ (ty1 --> ty2)) v2, substExp s (TAPP v1 v2) = TAPP (substVal s v1) (substVal s v2). auto. Qed.
|
||||
|
||||
End SubstDefs.
|
||||
|
||||
Hint Rewrite substTVAR substTPAIR substTINT substTBOOL substTFST substTSND substTFIX substTOP substTGT substTVAL substTLET substTIF substTAPP : substHints.
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Identity substitution
|
||||
==========================================================================*)
|
||||
|
||||
Definition idSubst env : Subst env env := fun ty (x:Var env ty) => TVAR x.
|
||||
Implicit Arguments idSubst [].
|
||||
|
||||
Lemma liftIdSubst : forall env ty, liftSubst ty (idSubst env) = idSubst (ty :: env).
|
||||
intros env ty. apply MapExtensional. unfold liftSubst. intros.
|
||||
dependent destruction var; auto.
|
||||
Qed.
|
||||
|
||||
Lemma applyIdSubst :
|
||||
(forall env ty (v : Value env ty), substVal (idSubst env) v = v)
|
||||
/\ (forall env ty (e : Exp env ty), substExp (idSubst env) e = e).
|
||||
Proof.
|
||||
apply ExpVal_ind; intros; autorewrite with substHints; try rewrite liftIdSubst; try rewrite liftIdSubst; try rewrite H; try rewrite H0; try rewrite H1; auto.
|
||||
Qed.
|
||||
|
||||
Notation "[ x , .. , y ]" := (consMap x .. (consMap y (idSubst _)) ..) : Subst_scope.
|
||||
Delimit Scope Subst_scope with subst.
|
||||
Arguments Scope substExp [_ _ Subst_scope].
|
||||
Arguments Scope substVal [_ _ Subst_scope].
|
||||
|
||||
Lemma LiftSubstDef : forall env env' ty (s : Subst env' env), liftSubst ty s = consMap (TVAR (ZVAR _ _)) (shiftSubst ty s).
|
||||
intros. apply MapExtensional. intros. dependent destruction var; auto. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Composition of substitution followed by renaming
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeRenamingSubst env env' env'' (r : Renaming env' env'') (s : Subst env env') : Subst env env'' := fun t var => renameVal r (s _ var).
|
||||
|
||||
Lemma liftComposeRenamingSubst : forall E env' env'' t (r:Renaming env' env'') (s:Subst E env'), liftSubst t (composeRenamingSubst r s) = composeRenamingSubst (liftRenaming t r) (liftSubst t s).
|
||||
intros. apply MapExtensional. intros t0 var. dependent induction var; auto.
|
||||
simpl. unfold composeRenamingSubst. unfold liftSubst. simpl. unfold renameVal at 1. rewrite <- (proj1 applyComposeRenaming). unfold renameVal. rewrite <- (proj1 applyComposeRenaming).
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma applyComposeRenamingSubst :
|
||||
(forall E t (v:Value E t) env' env'' (r:Renaming env' env'') (s : Subst E env'),
|
||||
substVal (composeRenamingSubst r s) v = renameVal r (substVal s v))
|
||||
/\ (forall E t (e:Exp E t) env' env'' (r:Renaming env' env'') (s : Subst E env'),
|
||||
substExp (composeRenamingSubst r s) e = renameExp r (substExp s e)).
|
||||
Proof.
|
||||
apply ExpVal_ind;
|
||||
(intros; try (autorewrite with substHints renameHints using try rewrite liftComposeRenamingSubst; try rewrite liftComposeRenamingSubst; try rewrite H; try rewrite H0; try rewrite H1); auto).
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Composition of substitutions
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeSubst env env' env'' (s' : Subst env' env'') (s : Subst env env') : Subst env env'' := fun ty var => substVal s' (s _ var).
|
||||
|
||||
Arguments Scope composeSubst [_ _ _ Subst_scope Subst_scope].
|
||||
|
||||
Lemma liftComposeSubst : forall env env' env'' ty (s' : Subst env' env'') (s : Subst env env'), liftSubst ty (composeSubst s' s) = composeSubst (liftSubst ty s') (liftSubst ty s).
|
||||
intros. apply MapExtensional. intros t0 var. dependent destruction var. auto.
|
||||
unfold composeSubst. simpl.
|
||||
rewrite <- (proj1 applyComposeRenamingSubst). unfold composeRenamingSubst. unfold substVal.
|
||||
rewrite <- (proj1 applyComposeRenaming). auto.
|
||||
Qed.
|
||||
|
||||
Lemma substComposeSubst :
|
||||
(forall env ty (v : Value env ty) env' env'' (s' : Subst env' env'') (s : Subst env env'),
|
||||
substVal (composeSubst s' s) v = substVal s' (substVal s v))
|
||||
/\ (forall env ty (e : Exp env ty) env' env'' (s' : Subst env' env'') (s : Subst env env'),
|
||||
substExp (composeSubst s' s) e = substExp s' (substExp s e)).
|
||||
Proof.
|
||||
apply ExpVal_ind;
|
||||
(intros; try (autorewrite with substHints using try rewrite liftComposeSubst; try rewrite liftComposeSubst; try rewrite H; try rewrite H0; try rewrite H1); auto).
|
||||
Qed.
|
||||
|
||||
Lemma composeCons : forall env env' env'' ty (s':Subst env' env'') (s:Subst env env') (v:Value _ ty),
|
||||
composeSubst (consMap v s') (liftSubst ty s) = consMap v (composeSubst s' s).
|
||||
intros. apply MapExtensional. intros t0 var. dependent destruction var. auto.
|
||||
unfold composeSubst. simpl. unfold substVal. rewrite <- (proj1 applyComposeRenaming). unfold composeRenaming. simpl.
|
||||
assert ((fun (t0:Ty) (var0:Var env' t0) => s' t0 var0) = s') by (apply MapExtensional; auto).
|
||||
rewrite H. auto.
|
||||
Qed.
|
||||
|
||||
Lemma composeSubstIdLeft : forall env env' (s : Subst env env'), composeSubst (idSubst _) s = s.
|
||||
Proof. intros. apply MapExtensional. intros. unfold composeSubst. apply (proj1 applyIdSubst). Qed.
|
||||
|
||||
Lemma composeSubstIdRight : forall env env' (s:Subst env env'), composeSubst s (idSubst _) = s.
|
||||
Proof. intros. apply MapExtensional. auto.
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Constructors are injective
|
||||
==========================================================================*)
|
||||
|
||||
Lemma TPAIR_injective :
|
||||
forall env ty1 ty2 (v1 : Value env ty1) (v2 : Value env ty2) v3 v4,
|
||||
TPAIR v1 v2 = TPAIR v3 v4 ->
|
||||
v1 = v3 /\ v2 = v4.
|
||||
intros. dependent destruction H. auto.
|
||||
Qed.
|
||||
|
||||
Lemma TFST_injective :
|
||||
forall env ty1 ty2 (v1 : Value env (ty1 ** ty2)) v2,
|
||||
TFST v1 = TFST v2 ->
|
||||
v1 = v2.
|
||||
intros. dependent destruction H. auto.
|
||||
Qed.
|
||||
|
||||
Lemma TSND_injective :
|
||||
forall env ty1 ty2 (v1 : Value env (ty1 ** ty2)) v2,
|
||||
TSND v1 = TSND v2 ->
|
||||
v1 = v2.
|
||||
intros. dependent destruction H. auto.
|
||||
Qed.
|
||||
|
||||
Lemma TVAL_injective :
|
||||
forall env ty (v1 : Value env ty) v2,
|
||||
TVAL v1 = TVAL v2 ->
|
||||
v1 = v2.
|
||||
intros. dependent destruction H. auto.
|
||||
Qed.
|
||||
|
||||
Lemma TFIX_injective :
|
||||
forall env ty1 ty2 (e1 : Exp (ty1 :: _ :: env) ty2) e2,
|
||||
TFIX e1 = TFIX e2 ->
|
||||
e1 = e2.
|
||||
intros. dependent destruction H. auto.
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Closed forms
|
||||
==========================================================================*)
|
||||
Lemma ClosedFunction : forall ty1 ty2 (v : CValue (ty1 --> ty2)), exists b, v = TFIX b.
|
||||
unfold CValue.
|
||||
intros ty1 ty2 v.
|
||||
dependent destruction v.
|
||||
(* TVAR case: impossible *)
|
||||
dependent destruction v.
|
||||
(* TFIX case *)
|
||||
exists e. trivial.
|
||||
Qed.
|
||||
|
||||
Lemma ClosedPair : forall ty1 ty2 (v : CValue (ty1 ** ty2)), exists v1, exists v2, v = TPAIR v1 v2.
|
||||
unfold CValue.
|
||||
intros ty1 ty2 v.
|
||||
dependent destruction v.
|
||||
(* TVAR case: impossible *)
|
||||
dependent destruction v.
|
||||
(* TPAIR case *)
|
||||
exists v1. exists v2. trivial.
|
||||
Qed.
|
||||
|
||||
Lemma ClosedInt : forall (v : CValue Int), exists i, v = TINT i.
|
||||
unfold CValue.
|
||||
intros v. dependent destruction v.
|
||||
exists n. trivial.
|
||||
dependent destruction v.
|
||||
Qed.
|
||||
|
||||
Lemma ClosedBool : forall (v : CValue Bool), exists b, v = TBOOL b.
|
||||
unfold CValue.
|
||||
intros v. dependent destruction v.
|
||||
exists b. trivial.
|
||||
dependent destruction v.
|
||||
Qed.
|
||||
|
||||
Unset Implicit Arguments.
|
67
papers/domains-2009/typedopsem.v
Executable file
67
papers/domains-2009/typedopsem.v
Executable file
|
@ -0,0 +1,67 @@
|
|||
(* typedopsem.v
|
||||
Standard operational semantics
|
||||
*)
|
||||
Require Import utility.
|
||||
Require Import Coq.Program.Equality.
|
||||
Require Import EqdepFacts.
|
||||
Require Import typedlambda.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
(*==========================================================================
|
||||
Evaluation relation
|
||||
==========================================================================*)
|
||||
|
||||
(** printing =>> %\ensuremath{\Downarrow}% *)
|
||||
(** printing n1 %\ensuremath{n_1}% *)
|
||||
(** printing n2 %\ensuremath{n_2}% *)
|
||||
(** printing v1 %\ensuremath{v_1}% *)
|
||||
(** printing v2 %\ensuremath{v_2}% *)
|
||||
(** printing e1 %\ensuremath{e_1}% *)
|
||||
(** printing e2 %\ensuremath{e_2}% *)
|
||||
Reserved Notation "e =>> v" (at level 70, no associativity).
|
||||
|
||||
Open Scope Subst_scope.
|
||||
Inductive Ev: forall ty, CExp ty -> CValue ty -> Prop :=
|
||||
| e_Val: forall ty (v : CValue ty), TVAL v =>> v
|
||||
| e_Op: forall op n1 n2, TOP op (TINT n1) (TINT n2) =>> TINT (op n1 n2)
|
||||
| e_Gt : forall n1 n2, TGT (TINT n1) (TINT n2) =>> TBOOL (ble_nat n2 n1)
|
||||
| e_Fst : forall ty1 ty2 (v1 : CValue ty1) (v2 : CValue ty2), TFST (TPAIR v1 v2) =>> v1
|
||||
| e_Snd : forall ty1 ty2 (v1 : CValue ty1) (v2 : CValue ty2), TSND (TPAIR v1 v2) =>> v2
|
||||
| e_App : forall ty1 ty2 e (v1 : CValue ty1) (v2 : CValue ty2), substExp [ v1, TFIX e ] e =>> v2 -> TAPP (TFIX e) v1 =>> v2
|
||||
| e_Let : forall ty1 ty2 e1 e2 (v1 : CValue ty1) (v2 : CValue ty2), e1 =>> v1 -> substExp [ v1 ] e2 =>> v2 -> TLET e1 e2 =>> v2
|
||||
| e_IfTrue : forall ty (e1 e2 : CExp ty) v, e1 =>> v -> TIF (TBOOL true) e1 e2 =>> v
|
||||
| e_IfFalse : forall ty (e1 e2 : CExp ty) v, e2 =>> v -> TIF (TBOOL false) e1 e2 =>> v
|
||||
where "e '=>>' v" := (Ev e v).
|
||||
|
||||
(*==========================================================================
|
||||
Determinacy of evaluation
|
||||
==========================================================================*)
|
||||
|
||||
Lemma Determinacy: forall ty, forall (e : CExp ty) v1, e =>> v1 -> forall v2, e =>> v2 -> v1 = v2.
|
||||
Proof.
|
||||
intros ty e v1 Ev1.
|
||||
induction Ev1.
|
||||
|
||||
(* e_Val *)
|
||||
intros v2 Ev2. dependent destruction Ev2; trivial.
|
||||
(* e_Op *)
|
||||
intros v2 Ev2. dependent destruction Ev2; trivial.
|
||||
(* e_Gt *)
|
||||
intros v2 Ev2. dependent destruction Ev2; trivial.
|
||||
(* e_Fst *)
|
||||
intros v3 Ev3. dependent destruction Ev3; trivial.
|
||||
(* e_Snd *)
|
||||
intros v3 Ev3. dependent destruction Ev3; trivial.
|
||||
(* e_App *)
|
||||
intros v3 Ev3. dependent destruction Ev3; auto.
|
||||
(* e_Let *)
|
||||
intros v3 Ev3. dependent destruction Ev3. rewrite (IHEv1_1 _ Ev3_1) in *. auto.
|
||||
(* e_IfTrue *)
|
||||
intros v3 Ev3. dependent destruction Ev3; auto.
|
||||
(* e_IfFalse *)
|
||||
intros v3 Ev3. dependent destruction Ev3; auto.
|
||||
Qed.
|
||||
|
||||
Unset Implicit Arguments.
|
52
papers/domains-2009/typedsoundness.v
Executable file
52
papers/domains-2009/typedsoundness.v
Executable file
|
@ -0,0 +1,52 @@
|
|||
Require Import utility.
|
||||
Require Import PredomAll.
|
||||
Require Import typeddensem.
|
||||
Require Import typedopsem.
|
||||
Require Import typedsubst.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
Theorem Soundness: forall ty (e : CExp ty) v, e =>> v -> SemExp e == eta << SemVal v.
|
||||
Proof.
|
||||
intros t e v Ev.
|
||||
induction Ev.
|
||||
|
||||
(* e_Val *)
|
||||
simpl. trivial.
|
||||
|
||||
(* e_Op *)
|
||||
simpl. rewrite fcont_comp_assoc. apply fcont_eq_intro; auto.
|
||||
|
||||
(* e_Gt *)
|
||||
simpl. apply fcont_eq_intro; auto.
|
||||
|
||||
(* e_Fst *)
|
||||
simpl. apply fcont_eq_intro; auto.
|
||||
|
||||
(* e_Snd *)
|
||||
simpl. apply fcont_eq_intro; auto.
|
||||
|
||||
(* e_App *)
|
||||
rewrite <- IHEv. clear IHEv. clear Ev.
|
||||
|
||||
rewrite <- (proj2 SemCommutesWithSubst _ _ e nil (consMap v1 (consMap (TFIX e) (idSubst _)))).
|
||||
simpl.
|
||||
assert (H := DOne_final (D := DOne) ID (K DOne (tt:DOne))). rewrite <- H.
|
||||
rewrite hdConsMap.
|
||||
rewrite FIXP_com at 1. refine (fcont_eq_intro _) ; auto.
|
||||
|
||||
(* e_Let *)
|
||||
simpl. rewrite <- IHEv2. clear IHEv2.
|
||||
rewrite <- (proj2 SemCommutesWithSubst _ _ e2 _ (consMap v1 (idSubst _))).
|
||||
simpl. rewrite <- (KLEISLIR_unit (SemExp e2) (SemVal v1) (K DOne (tt:DOne))).
|
||||
rewrite <- IHEv1. clear IHEv1.
|
||||
rewrite (DOne_final (D := DOne) ID (K DOne (tt:DOne))).
|
||||
trivial.
|
||||
|
||||
(* e_Iftrue *)
|
||||
simpl. apply fcont_eq_intro; auto.
|
||||
|
||||
(* e_Iffalse *)
|
||||
simpl. apply fcont_eq_intro; auto.
|
||||
Qed.
|
298
papers/domains-2009/typedsubst.v
Executable file
298
papers/domains-2009/typedsubst.v
Executable file
|
@ -0,0 +1,298 @@
|
|||
Require Import utility.
|
||||
Require Import PredomAll.
|
||||
Require Import typedlambda.
|
||||
Require Import typeddensem.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
(* Need this for nice pretty-printing *)
|
||||
Unset Printing Implicit Defensive.
|
||||
|
||||
(*==========================================================================
|
||||
Semantics of substitution and renaming
|
||||
==========================================================================*)
|
||||
|
||||
(* Apply semantic function pointwise to a renaming or substitution *)
|
||||
Fixpoint SemSubst env env' : Subst env' env -> SemEnv env -c> SemEnv env' :=
|
||||
match env' with
|
||||
| nil => fun s => K _ (tt : DOne)
|
||||
| _ :: _ => fun s => <| SemSubst (tlMap s) , SemVal (hdMap s) |>
|
||||
end.
|
||||
|
||||
Fixpoint SemRenaming env env' : Renaming env' env -> SemEnv env -c> SemEnv env' :=
|
||||
match env' with
|
||||
| nil => fun s => K _ (tt : DOne)
|
||||
| _ :: _ => fun s => <| SemRenaming (tlMap s) , SemVar (hdMap s) |>
|
||||
end.
|
||||
|
||||
Lemma SemLiftRenaming :
|
||||
forall env env' (r : Renaming env env') ty, SemRenaming (tlMap (liftRenaming ty r)) == SemRenaming r << FST.
|
||||
Proof.
|
||||
induction env. intros. simpl.
|
||||
apply (DOne_final (K (Dprod (SemEnv env') (SemTy ty)) (tt:DOne)) _).
|
||||
|
||||
intros.
|
||||
simpl. destruct (consMapInv r) as [r' [var eq]].
|
||||
subst. simpl. rewrite hdConsMap. rewrite tlConsMap.
|
||||
specialize (IHenv env' r' ty).
|
||||
rewrite PROD_fun_comp.
|
||||
refine (PROD_fun_eq_compat _ _).
|
||||
rewrite IHenv. auto. auto.
|
||||
Qed.
|
||||
|
||||
Lemma SemLift2Renaming :
|
||||
forall env env' (r : Renaming env env') ty ty', SemRenaming (tlMap (tlMap (liftRenaming ty' (liftRenaming ty r)))) == SemRenaming r << FST << FST.
|
||||
Proof.
|
||||
induction env. intros. simpl.
|
||||
apply (DOne_final (K (Dprod (Dprod (SemEnv env') (SemTy ty)) (SemTy ty')) (tt:DOne)) _).
|
||||
|
||||
intros.
|
||||
simpl. destruct (consMapInv r) as [r' [var eq]].
|
||||
subst. simpl. rewrite hdConsMap. rewrite tlConsMap.
|
||||
specialize (IHenv env' r' ty).
|
||||
rewrite PROD_fun_comp. rewrite PROD_fun_comp.
|
||||
refine (PROD_fun_eq_compat _ _).
|
||||
rewrite IHenv. auto. auto.
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Semantic function commutes with renaming
|
||||
==========================================================================*)
|
||||
|
||||
Lemma SemCommutesWithRenaming:
|
||||
(forall env ty (v : Value env ty) env' (r : Renaming env env'),
|
||||
SemVal v << SemRenaming r == SemVal (renameVal r v))
|
||||
/\ (forall env ty (exp : Exp env ty) env' (r : Renaming env env'),
|
||||
SemExp exp << SemRenaming r == SemExp (renameExp r exp)).
|
||||
Proof.
|
||||
apply ExpVal_ind.
|
||||
|
||||
(* TINT *)
|
||||
intros. simpl. rewrite <- K_com; trivial.
|
||||
|
||||
(* TBOOL *)
|
||||
intros. simpl. rewrite <- K_com; trivial.
|
||||
|
||||
(* TVAR *)
|
||||
intros env ty var env' r.
|
||||
simpl.
|
||||
induction var.
|
||||
simpl. rewrite SND_PROD_fun.
|
||||
destruct (consMapInv r) as [r' [v eq]]. subst.
|
||||
simpl. rewrite hdConsMap. trivial.
|
||||
destruct (consMapInv r) as [r' [v eq]]. subst.
|
||||
simpl.
|
||||
specialize (IHvar r').
|
||||
rewrite <- IHvar. rewrite fcont_comp_assoc.
|
||||
rewrite tlConsMap. rewrite FST_PROD_fun. trivial.
|
||||
|
||||
(* TFIX *)
|
||||
intros E t t1 body IH E' s.
|
||||
specialize (IH _ (liftRenaming _ (liftRenaming _ s))).
|
||||
rewrite renameTFIX.
|
||||
simpl SemVal.
|
||||
rewrite fcont_comp_assoc.
|
||||
rewrite Curry_comp. rewrite Curry_comp. rewrite <- IH.
|
||||
clear IH. simpl. unfold PROD_map. rewrite fcont_comp_unitL. rewrite fcont_comp_unitL.
|
||||
rewrite PROD_fun_comp. rewrite SemLift2Renaming. trivial.
|
||||
|
||||
(* TPAIR *)
|
||||
intros E t1 t2 v1 IH1 v2 IH2 E' s.
|
||||
rewrite renameTPAIR. simpl SemVal. rewrite <- IH1. rewrite <- IH2. rewrite <- PROD_fun_comp. trivial.
|
||||
|
||||
(* TFST *)
|
||||
intros E t1 t2 v IH E' s. rewrite renameTFST. simpl.
|
||||
rewrite <- IH. repeat (rewrite fcont_comp_assoc). trivial.
|
||||
|
||||
(* TSND *)
|
||||
intros E t1 t2 v IH E' s. rewrite renameTSND. simpl.
|
||||
rewrite <- IH. repeat (rewrite fcont_comp_assoc). trivial.
|
||||
|
||||
(* TOP *)
|
||||
intros E n v1 IH1 v2 IH2 E' s. rewrite renameTOP. simpl.
|
||||
repeat (rewrite fcont_comp_assoc). rewrite <- IH1. rewrite <- IH2.
|
||||
rewrite PROD_fun_comp. trivial.
|
||||
|
||||
(* TGT *)
|
||||
intros E v1 IH1 v2 IH2 E' s. rewrite renameTGT. simpl.
|
||||
repeat (rewrite fcont_comp_assoc). rewrite <- IH1. rewrite <- IH2.
|
||||
rewrite PROD_fun_comp. trivial.
|
||||
|
||||
(* TVAL *)
|
||||
intros E t v IH E' s. rewrite renameTVAL. simpl.
|
||||
rewrite <- IH. rewrite fcont_comp_assoc. trivial.
|
||||
|
||||
(* TLET *)
|
||||
intros E t1 t2 e2 IH2 e1 IH1 E' s. rewrite renameTLET. simpl.
|
||||
rewrite fcont_comp_assoc.
|
||||
specialize (IH2 _ s).
|
||||
specialize (IH1 _ (liftRenaming _ s)).
|
||||
rewrite PROD_fun_comp.
|
||||
rewrite KLEISLIR_comp.
|
||||
rewrite <- IH2. clear IH2. rewrite <- IH1. clear IH1. simpl.
|
||||
unfold PROD_map. rewrite fcont_comp_unitL. rewrite fcont_comp_unitL.
|
||||
rewrite SemLiftRenaming. trivial.
|
||||
|
||||
(* TAPP *)
|
||||
intros E t1 t2 v1 IH1 v2 IH2 E' s. rewrite renameTAPP. simpl.
|
||||
rewrite fcont_comp_assoc. rewrite <- IH1. rewrite <- IH2.
|
||||
rewrite PROD_fun_comp. trivial.
|
||||
|
||||
(* TIF *)
|
||||
intros E t ec IHc te1 IH1 te2 IH2 E' s. rewrite renameTIF. simpl.
|
||||
rewrite fcont_comp3_comp. rewrite <- IHc. rewrite <- IH1. rewrite <- IH2. trivial.
|
||||
Qed.
|
||||
|
||||
Set Printing Implicit Defensive.
|
||||
|
||||
Lemma SemShiftRenaming :
|
||||
forall env env' (r : Renaming env env') ty, SemRenaming (shiftRenaming ty r) == SemRenaming r << FST.
|
||||
Proof.
|
||||
induction env. intros. simpl.
|
||||
apply (DOne_final (K (Dprod (SemEnv env') (SemTy ty)) (tt:DOne)) _).
|
||||
|
||||
intros.
|
||||
simpl. destruct (consMapInv r) as [r' [var eq]].
|
||||
subst. simpl. rewrite hdConsMap. rewrite tlConsMap.
|
||||
specialize (IHenv env' r' ty).
|
||||
rewrite PROD_fun_comp.
|
||||
refine (PROD_fun_eq_compat IHenv _); auto.
|
||||
Qed.
|
||||
|
||||
Lemma SemIdRenaming : forall env, SemRenaming (idRenaming env) == ID.
|
||||
Proof.
|
||||
induction env.
|
||||
simpl.
|
||||
apply (DOne_final (K DOne _) _).
|
||||
|
||||
rewrite idRenamingDef. simpl SemRenaming. rewrite tlConsMap. rewrite SemShiftRenaming.
|
||||
rewrite IHenv. rewrite fcont_comp_unitL.
|
||||
apply fcont_eq_intro. intros. destruct x. auto.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma SemShiftSubst :
|
||||
forall env env' (s : Subst env env') ty, SemSubst (shiftSubst ty s) == SemSubst s << FST.
|
||||
Proof.
|
||||
induction env. intros. simpl.
|
||||
apply (DOne_final (K (Dprod (SemEnv env') (SemTy ty)) (tt:DOne)) _).
|
||||
|
||||
intros.
|
||||
simpl. destruct (consMapInv s) as [s' [var eq]].
|
||||
subst. simpl. rewrite hdConsMap. rewrite tlConsMap.
|
||||
specialize (IHenv env' s' ty).
|
||||
rewrite PROD_fun_comp.
|
||||
refine (PROD_fun_eq_compat _ _).
|
||||
rewrite IHenv. auto. unfold shiftSubst.
|
||||
rewrite shiftConsMap. rewrite hdConsMap.
|
||||
simpl. rewrite <- (proj1 SemCommutesWithRenaming).
|
||||
refine (fcont_comp_eq_compat _ _). trivial.
|
||||
(* Seems a bit round-about *)
|
||||
assert ((fun t0 => SVAR ty) = tlMap (liftRenaming ty (idRenaming env'))).
|
||||
rewrite LiftRenamingDef. rewrite tlConsMap. apply MapExtensional. auto.
|
||||
rewrite H.
|
||||
rewrite SemLiftRenaming. rewrite SemIdRenaming. rewrite fcont_comp_unitL. auto.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma SemLiftSubst :
|
||||
forall env env' (s : Subst env env') ty, SemSubst (tlMap (liftSubst ty s)) == SemSubst s << FST.
|
||||
Proof.
|
||||
intros.
|
||||
rewrite LiftSubstDef. rewrite tlConsMap. apply SemShiftSubst.
|
||||
Qed.
|
||||
|
||||
Lemma SemLift2Subst :
|
||||
forall env env' (s : Subst env env') ty ty', SemSubst (tlMap (tlMap (liftSubst ty' (liftSubst ty s)))) == SemSubst s << FST << FST.
|
||||
Proof.
|
||||
intros.
|
||||
rewrite LiftSubstDef. rewrite tlConsMap. rewrite LiftSubstDef. unfold shiftSubst. rewrite shiftConsMap. rewrite tlConsMap.
|
||||
rewrite SemShiftSubst. rewrite SemShiftSubst. auto.
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Semantic function commutes with substitution
|
||||
==========================================================================*)
|
||||
|
||||
|
||||
Lemma SemCommutesWithSubst:
|
||||
(forall env ty (v : Value env ty) env' (s : Subst env env'),
|
||||
SemVal v << SemSubst s == SemVal (substVal s v))
|
||||
/\ (forall env ty (e : Exp env ty) env' (s : Subst env env'),
|
||||
SemExp e << SemSubst s == SemExp (substExp s e)).
|
||||
Proof.
|
||||
apply ExpVal_ind.
|
||||
|
||||
(* TINT *)
|
||||
intros. simpl. rewrite <- K_com; trivial.
|
||||
|
||||
(* TBOOL *)
|
||||
intros. simpl. rewrite <- K_com; trivial.
|
||||
|
||||
(* TVAR *)
|
||||
intros env ty var env' s.
|
||||
simpl.
|
||||
unfold substVal. simpl travVal.
|
||||
induction var. simpl. rewrite SND_PROD_fun.
|
||||
destruct (consMapInv s) as [s' [v eq]]. subst.
|
||||
simpl. rewrite hdConsMap. trivial.
|
||||
destruct (consMapInv s) as [s' [v eq]]. subst.
|
||||
simpl.
|
||||
specialize (IHvar s').
|
||||
rewrite <- IHvar. rewrite fcont_comp_assoc.
|
||||
rewrite tlConsMap. rewrite FST_PROD_fun. trivial.
|
||||
|
||||
(* TFIX *)
|
||||
intros E t t1 body IH E' s. rewrite substTFIX. simpl.
|
||||
rewrite fcont_comp_assoc.
|
||||
rewrite Curry_comp. rewrite Curry_comp. rewrite <- IH.
|
||||
clear IH. simpl. unfold PROD_map. rewrite fcont_comp_unitL. rewrite fcont_comp_unitL.
|
||||
rewrite SemLift2Subst. rewrite fcont_comp_assoc. rewrite PROD_fun_comp. rewrite fcont_comp_assoc. trivial.
|
||||
|
||||
(* TPAIR *)
|
||||
intros E t1 t2 v1 IH1 v2 IH2 E' s. rewrite substTPAIR. simpl.
|
||||
rewrite <- IH1. rewrite <- IH2. rewrite PROD_fun_comp. trivial.
|
||||
|
||||
(* TFST *)
|
||||
intros E t1 t2 v IH E' s. rewrite substTFST. simpl.
|
||||
rewrite <- IH. repeat (rewrite fcont_comp_assoc). trivial.
|
||||
|
||||
(* TSND *)
|
||||
intros E t1 t2 v IH E' s. rewrite substTSND. simpl.
|
||||
rewrite <- IH. repeat (rewrite fcont_comp_assoc). trivial.
|
||||
|
||||
(* TOP *)
|
||||
intros E n v1 IH1 v2 IH2 E' s. rewrite substTOP. simpl.
|
||||
repeat (rewrite fcont_comp_assoc). rewrite <- IH1. rewrite <- IH2.
|
||||
rewrite PROD_fun_comp. trivial.
|
||||
|
||||
(* TGT *)
|
||||
intros E v1 IH1 v2 IH2 E' s. rewrite substTGT. simpl.
|
||||
repeat (rewrite fcont_comp_assoc). rewrite <- IH1. rewrite <- IH2.
|
||||
rewrite PROD_fun_comp. trivial.
|
||||
|
||||
(* TVAL *)
|
||||
intros E t v IH E' s. rewrite substTVAL. simpl.
|
||||
rewrite <- IH. rewrite fcont_comp_assoc. trivial.
|
||||
|
||||
(* TLET *)
|
||||
intros E t1 t2 e2 IH2 e1 IH1 E' s. rewrite substTLET. simpl.
|
||||
rewrite fcont_comp_assoc.
|
||||
specialize (IH2 _ s).
|
||||
specialize (IH1 _ (liftSubst _ s)).
|
||||
rewrite PROD_fun_comp.
|
||||
rewrite KLEISLIR_comp.
|
||||
rewrite <- IH2. clear IH2. rewrite <- IH1. clear IH1. simpl.
|
||||
unfold PROD_map. rewrite fcont_comp_unitL. rewrite fcont_comp_unitL. rewrite SemLiftSubst. trivial.
|
||||
|
||||
(* TAPP *)
|
||||
intros E t1 t2 v1 IH1 v2 IH2 E' s. rewrite substTAPP. simpl.
|
||||
rewrite fcont_comp_assoc. rewrite <- IH1. rewrite <- IH2.
|
||||
rewrite PROD_fun_comp. trivial.
|
||||
|
||||
(* TIF *)
|
||||
intros E t ec IHc te1 IH1 te2 IH2 E' s. rewrite substTIF. simpl.
|
||||
rewrite fcont_comp3_comp.
|
||||
rewrite <- IHc. rewrite <- IH1. rewrite <- IH2. trivial.
|
||||
Qed.
|
256
papers/domains-2009/uni.v
Executable file
256
papers/domains-2009/uni.v
Executable file
|
@ -0,0 +1,256 @@
|
|||
Require Export utility.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
Inductive Value : Set :=
|
||||
| VAR : nat -> Value
|
||||
| NUM : nat -> Value
|
||||
| LAMBDA : Exp -> Value
|
||||
with Exp : Set :=
|
||||
| VAL : Value -> Exp
|
||||
| APP : Value -> Value -> Exp
|
||||
| LET : Exp -> Exp -> Exp
|
||||
| IFZ : Value -> Exp -> Exp -> Exp
|
||||
| OP : (nat -> nat -> nat) -> Value -> Value -> Exp.
|
||||
|
||||
Scheme Value_rec2 := Induction for Value Sort Set
|
||||
with Exp_rec2 := Induction for Exp Sort Set.
|
||||
|
||||
Scheme Value_type2 := Induction for Value Sort Type
|
||||
with Exp_type2 := Induction for Exp Sort Type.
|
||||
|
||||
Scheme Value_ind2 := Induction for Value Sort Prop
|
||||
with Exp_ind2 := Induction for Exp Sort Prop.
|
||||
|
||||
Combined Scheme ExpValue_ind from Value_ind2, Exp_ind2.
|
||||
|
||||
Inductive VTyping (n:nat) : Value -> Type :=
|
||||
| TVAR : forall m, m < n -> VTyping n (VAR m)
|
||||
| TNUM : forall m, VTyping n (NUM m)
|
||||
| TLAMBDA : forall e, ETyping (S n) e -> VTyping n (LAMBDA e)
|
||||
with ETyping (n:nat) : Exp -> Type :=
|
||||
| TVAL : forall v, VTyping n v -> ETyping n (VAL v)
|
||||
| TAPP : forall e1 e2, VTyping n e1 -> VTyping n e2 -> ETyping n (APP e1 e2)
|
||||
| TLET : forall e1 e2, ETyping n e1 -> ETyping (S n) e2 -> ETyping n (LET e1 e2)
|
||||
| TIFZ : forall v e1 e2, VTyping n v -> ETyping n e1 -> ETyping n e2 -> ETyping n (IFZ v e1 e2)
|
||||
| TOP : forall op v1 v2, VTyping n v1 -> VTyping n v2 -> ETyping n (OP op v1 v2).
|
||||
|
||||
Hint Resolve TVAR TNUM TLAMBDA TVAL TAPP TLET TOP TIFZ.
|
||||
|
||||
Scheme VTyping_rec2 := Induction for VTyping Sort Set
|
||||
with ETyping_rec2 := Induction for ETyping Sort Set.
|
||||
|
||||
Scheme VTyping_rect2 := Induction for VTyping Sort Type
|
||||
with ETyping_rect2 := Induction for ETyping Sort Type.
|
||||
|
||||
Scheme VTyping_ind2 := Induction for VTyping Sort Prop
|
||||
with ETyping_ind2 := Induction for ETyping Sort Prop.
|
||||
|
||||
Combined Scheme Typing_ind from VTyping_ind2, ETyping_ind2.
|
||||
|
||||
Fixpoint shiftV (k n:nat) (e:Value) {struct e} : Value :=
|
||||
match e with
|
||||
| VAR m => (if bleq_nat k m then VAR (m + n) else VAR m)
|
||||
| NUM m => NUM m
|
||||
| LAMBDA e => LAMBDA (shiftE (S k) n e)
|
||||
end
|
||||
with shiftE (k n:nat) (e:Exp) {struct e} : Exp :=
|
||||
match e with
|
||||
| VAL v => VAL (shiftV k n v)
|
||||
| APP e1 e2 => APP (shiftV k n e1) (shiftV k n e2)
|
||||
| LET e1 e2 => LET (shiftE k n e1) (shiftE (S k) n e2)
|
||||
| IFZ v e1 e2 => IFZ (shiftV k n v) (shiftE k n e1) (shiftE k n e2)
|
||||
| OP op v1 v2 => OP op (shiftV k n v1) (shiftV k n v2)
|
||||
end.
|
||||
|
||||
Lemma WeakeningV n v (tv:VTyping n v) x k : VTyping (n + x) (shiftV k x v).
|
||||
intros n v tv x.
|
||||
apply (@VTyping_rect2 (fun m v tv => forall k, VTyping (m + x) (shiftV k x v))
|
||||
(fun m e te => forall k, ETyping (m + x) (shiftE k x e))) ;
|
||||
try (intros ; simpl ; auto ; fail).
|
||||
|
||||
intros y m ym k. simpl.
|
||||
case (bleq_nat k m).
|
||||
eapply TVAR. omega.
|
||||
eapply TVAR. omega.
|
||||
Defined.
|
||||
|
||||
Lemma WeakeningE n e (te:ETyping n e) x k : ETyping (n + x) (shiftE k x e).
|
||||
intros n e te x.
|
||||
apply (@ETyping_rect (fun m e te => forall k, ETyping (m + x) (shiftE k x e))) ;
|
||||
try (intros ; simpl ; auto ; fail).
|
||||
|
||||
intros m v tv k.
|
||||
simpl. eapply TVAL. eapply WeakeningV. apply tv.
|
||||
|
||||
intros m v1 v2 tv1 tv2 k. simpl. eapply TAPP. eapply WeakeningV. apply tv1.
|
||||
eapply WeakeningV. apply tv2.
|
||||
|
||||
intros m v e1 e2 tv te1 IH1 te2 IH2 k.
|
||||
simpl. eapply TIFZ. eapply WeakeningV. apply tv. apply IH1. apply IH2.
|
||||
|
||||
intros m op v1 v2 tv1 tv2 k. simpl.
|
||||
eapply TOP ; eapply WeakeningV ; auto.
|
||||
Defined.
|
||||
|
||||
Fixpoint substVal (vl:list Value) (e:Value) : Value :=
|
||||
match e with
|
||||
| VAR m => (match nth_error vl m with | value ee => ee | error => VAR m end)
|
||||
| NUM m => NUM m
|
||||
| LAMBDA e => LAMBDA (substExp (VAR 0 :: (map (shiftV 0 1) vl)) e)
|
||||
end
|
||||
with substExp (vl:list Value) (e:Exp) : Exp :=
|
||||
match e with
|
||||
| VAL v => VAL (substVal vl v)
|
||||
| APP e1 e2 => APP (substVal vl e1) (substVal vl e2)
|
||||
| LET e1 e2 => LET (substExp vl e1) (substExp (VAR 0 :: map (shiftV 0 1) vl) e2)
|
||||
| IFZ v e1 e2 => IFZ (substVal vl v) (substExp vl e1) (substExp vl e2)
|
||||
| OP op v1 v2 => OP op (substVal vl v1) (substVal vl v2)
|
||||
end.
|
||||
|
||||
Lemma subst_typingV:
|
||||
forall n v (tv:VTyping n v) n' vs, n = length vs ->
|
||||
(forall i v', nth_error vs i = Some v' -> VTyping n' v') ->
|
||||
VTyping n' (substVal vs v).
|
||||
Proof.
|
||||
apply (@VTyping_rect2
|
||||
(fun n v dv => forall n' vs, n = length vs ->
|
||||
(forall i v', nth_error vs i = Some v' -> VTyping n' v') ->
|
||||
VTyping n' (substVal vs v))
|
||||
(fun n e de => forall n' vs, n = length vs ->
|
||||
(forall i v', nth_error vs i = Some v' -> VTyping n' v') ->
|
||||
ETyping n' (substExp vs e))) ; intros n.
|
||||
|
||||
intros m mn n' vs nl C.
|
||||
simpl. case_eq (nth_error vs m).
|
||||
intros v nth.
|
||||
specialize (C _ _ nth).
|
||||
apply C.
|
||||
intros nth.
|
||||
assert False as F. clear C n'.
|
||||
generalize dependent m.
|
||||
generalize dependent n.
|
||||
induction vs. intros n nl m mn.
|
||||
rewrite nl in mn. simpl in mn. intros. omega.
|
||||
intros n nl m. simpl in nl. specialize (IHvs (length vs) (refl_equal _)).
|
||||
case m. intros _ incon. simpl in incon. inversion incon.
|
||||
clear m. intros m sm. simpl. rewrite nl in sm.
|
||||
specialize (IHvs _ (lt_S_n _ _ sm)). auto.
|
||||
inversion F.
|
||||
|
||||
simpl. intros. eapply TNUM.
|
||||
simpl.
|
||||
intros e te IH1 n' vs nl C.
|
||||
eapply TLAMBDA. apply IH1.
|
||||
simpl. rewrite map_length. auto.
|
||||
intros i v'. case i ; clear i.
|
||||
simpl. intros va. inversion va. eapply TVAR. omega.
|
||||
intros i. simpl. rewrite nth_error_map.
|
||||
case_eq (nth_error vs i). intros v nth va.
|
||||
inversion va. clear H0 va.
|
||||
simpl.
|
||||
assert (S n' = n' + 1) as A by omega. rewrite A.
|
||||
eapply (WeakeningV). eapply (C _ _ nth).
|
||||
intros _ incon. inversion incon.
|
||||
|
||||
intros v tv IH n' vs nl C. simpl.
|
||||
eapply (TVAL). apply IH ; auto.
|
||||
|
||||
intros v1 v2 tv1 IH1 tv2 IH2 n' vs nl C.
|
||||
simpl. eapply (TAPP). eapply IH1 ; auto. eapply IH2 ; auto.
|
||||
|
||||
intros e1 e2 te1 IH1 te2 IH2 n' vs nl C.
|
||||
simpl. eapply TLET. eapply IH1 ; auto.
|
||||
eapply IH2 ; simpl ; auto. rewrite map_length. auto.
|
||||
intros i v'. case i. simpl. intros va. inversion va.
|
||||
eapply TVAR. omega.
|
||||
clear i. intros i. simpl.
|
||||
rewrite nth_error_map. case_eq (nth_error vs i).
|
||||
intros v nth va. inversion va. clear H0 va.
|
||||
assert (S n' = n' + 1) as A by omega. rewrite A.
|
||||
eapply (WeakeningV).
|
||||
apply (C _ _ nth).
|
||||
intros _ incon. inversion incon.
|
||||
|
||||
intros v e1 e2 tv IHv te1 IH1 te2 IH2 n' vs nl C. simpl.
|
||||
eapply TIFZ. apply IHv ; auto. apply IH1 ; auto. apply IH2 ; auto.
|
||||
|
||||
intros op v1 v2 tv1 IH1 tv2 IH2 n' vs nl C. simpl.
|
||||
eapply TOP. apply IH1 ; auto. apply IH2 ; auto.
|
||||
Defined.
|
||||
|
||||
Lemma subst_typingE:
|
||||
forall n e (te:ETyping n e) n' vs, n = length vs ->
|
||||
(forall i v', nth_error vs i = Some v' -> VTyping n' v') ->
|
||||
ETyping n' (substExp vs e).
|
||||
Proof.
|
||||
eapply (@ETyping_rect (fun n e _ => forall n' vs, n = length vs ->
|
||||
(forall i v', nth_error vs i = Some v' -> VTyping n' v') ->
|
||||
ETyping n' (substExp vs e))).
|
||||
|
||||
intros n v tv n' vs nl C. simpl. eapply TVAL.
|
||||
apply (subst_typingV tv) ; auto.
|
||||
|
||||
intros n v1 v2 tv1 tv2 n' vs nl C.
|
||||
simpl. eapply TAPP. apply (subst_typingV tv1) ; auto.
|
||||
apply (subst_typingV tv2) ; auto.
|
||||
|
||||
intros n e1 e2 te1 IH1 te2 IH2 n' vs nl C.
|
||||
simpl. eapply TLET ; auto.
|
||||
apply IH2. simpl. rewrite map_length. auto.
|
||||
intros i v'. case i. simpl. intros va. inversion va. eapply TVAR ; auto. omega.
|
||||
clear i ; intros i. simpl. rewrite nth_error_map.
|
||||
case_eq (nth_error vs i). intros v nth va. inversion va. clear H0 va.
|
||||
assert (S n' = n' + 1) as A by omega. rewrite A.
|
||||
eapply (WeakeningV). apply (C _ _ nth).
|
||||
intros nth incon. inversion incon.
|
||||
|
||||
intros n v e1 e2 tv te1 IH1 te2 IH2 n' vs nl C. simpl.
|
||||
eapply TIFZ.
|
||||
eapply (subst_typingV (n:=n)) ; auto. apply IH1 ; auto. apply IH2 ; auto.
|
||||
|
||||
intros n op v1 v2 tv1 tv2 n' vs nl C. simpl.
|
||||
eapply TOP ; eapply (subst_typingV (n:=n)) ; auto.
|
||||
Defined.
|
||||
|
||||
Reserved Notation "e '=>>' v" (at level 75).
|
||||
|
||||
Inductive Evaluation : Exp -> Value -> Type :=
|
||||
| e_Value : forall v, VAL v =>> v
|
||||
| e_App : forall e1 v2 v,
|
||||
substExp [v2] e1 =>> v ->
|
||||
(APP (LAMBDA e1) v2) =>> v
|
||||
| e_Let : forall e1 v1 e2 v2, e1 =>> v1 -> substExp [v1] e2 =>> v2 ->
|
||||
(LET e1 e2) =>> v2
|
||||
| e_Ifz1 : forall e1 e2 v1, e1 =>> v1 -> IFZ (NUM 0) e1 e2 =>> v1
|
||||
| e_Ifz2 : forall e1 e2 v2 n, e2 =>> v2 -> IFZ (NUM (S n)) e1 e2 =>> v2
|
||||
| e_Op : forall op n1 n2, OP op (NUM n1) (NUM n2) =>> NUM (op n1 n2)
|
||||
where "e =>> v" := (Evaluation e v).
|
||||
|
||||
Lemma eval_preserve_type: forall e v, e =>> v -> ETyping 0 e -> VTyping 0 v.
|
||||
intros e v ev.
|
||||
induction ev ; intros ty ; auto.
|
||||
inversion ty. auto.
|
||||
|
||||
apply IHev. eapply (subst_typingE (n:=1)) ; auto.
|
||||
inversion ty. inversion H1. auto.
|
||||
|
||||
intros i. case i ; clear i. simpl. intros v' va. inversion va.
|
||||
rewrite <- H0.
|
||||
inversion ty. auto.
|
||||
intros i v'. simpl. intro incon. destruct i ; simpl in incon ; inversion incon.
|
||||
|
||||
apply IHev2. inversion ty.
|
||||
clear e0 e3 H H0.
|
||||
eapply (subst_typingE (n:=1)) ; simpl ; auto.
|
||||
intros i. case i ; clear i. intros v. simpl. intros va. inversion va.
|
||||
rewrite <- H0. apply IHev1. apply H1.
|
||||
intros n v'. simpl. intros incon. destruct n ; simpl in * ; inversion incon.
|
||||
|
||||
inversion ty.
|
||||
apply IHev ; auto.
|
||||
|
||||
inversion ty. apply IHev ; auto.
|
||||
Qed.
|
||||
|
1354
papers/domains-2009/uniade.v
Executable file
1354
papers/domains-2009/uniade.v
Executable file
File diff suppressed because it is too large
Load diff
337
papers/domains-2009/unisem.v
Executable file
337
papers/domains-2009/unisem.v
Executable file
|
@ -0,0 +1,337 @@
|
|||
Require Import PredomAll.
|
||||
Require Export lc.
|
||||
Require Export uni.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
Definition SimpleOpm (A B : Type) (op : A -> B) : Discrete A -m> Discrete B.
|
||||
intros A B op.
|
||||
exists (op).
|
||||
unfold monotonic.
|
||||
intros.
|
||||
simpl in *.
|
||||
rewrite H. auto.
|
||||
Defined.
|
||||
|
||||
Definition SimpleOp A B (op:A -> B) :
|
||||
Discrete A -c> Discrete B.
|
||||
intros.
|
||||
exists (SimpleOpm op).
|
||||
simpl in *.
|
||||
unfold continuous.
|
||||
intros.
|
||||
simpl. auto.
|
||||
Defined.
|
||||
|
||||
(*
|
||||
|
||||
Definition SimpleOp2mm (A B C : Type) (op : A -> B -> C) : Discrete A -> Discrete B -m> Discrete C.
|
||||
intros.
|
||||
exists (op X).
|
||||
unfold monotonic.
|
||||
intros.
|
||||
simpl in *.
|
||||
rewrite H. auto.
|
||||
Defined.
|
||||
|
||||
Definition SimpleOp2c A B C (op:A -> B -> C) :
|
||||
Discrete A -> Discrete B -c> Discrete C.
|
||||
intros.
|
||||
exists (SimpleOp2mm op X).
|
||||
simpl in *.
|
||||
unfold continuous.
|
||||
intros.
|
||||
simpl. auto.
|
||||
Defined.
|
||||
|
||||
Definition SimpleOp2m (A B C:Type) (op : A -> B -> C) :
|
||||
Discrete A -m> Discrete B -C-> Discrete C.
|
||||
intros.
|
||||
exists (SimpleOp2c op).
|
||||
unfold monotonic.
|
||||
intros.
|
||||
simpl in *.
|
||||
rewrite H. auto.
|
||||
Defined.
|
||||
|
||||
Definition SimpleOp2 A B C (op:A -> B -> C) :
|
||||
Discrete A -c> Discrete B -C-> Discrete C.
|
||||
intros.
|
||||
exists (SimpleOp2m op).
|
||||
unfold continuous ; intros ; simpl ; auto.
|
||||
Defined.
|
||||
*)
|
||||
|
||||
Definition FS := BiLift_strict (BiSum (BiConst (Discrete nat)) BiArrow).
|
||||
Definition DInf := DInf FS.
|
||||
Definition VInf := Dsum (Discrete nat) (DInf -C-> DInf).
|
||||
|
||||
Definition Roll : DL VInf -C-> DInf := ROLL FS.
|
||||
Definition Unroll : DInf -C-> DL VInf := UNROLL FS.
|
||||
|
||||
Lemma UR_eq : forall x, Unroll (Roll x) == x.
|
||||
intros y. rewrite <- fcont_comp_simpl.
|
||||
apply Oeq_trans with (y:= ID y).
|
||||
refine (app_eq_compat _ (Oeq_refl _)). unfold Unroll. unfold Roll. apply (DIso_ur FS). auto.
|
||||
Qed.
|
||||
|
||||
Lemma RU_eq : forall x, Roll (Unroll x) == x.
|
||||
intros x.
|
||||
assert (xx := fcont_eq_elim (DIso_ru FS) x). rewrite fcont_comp_simpl in xx. rewrite ID_simpl in xx.
|
||||
simpl in xx. unfold Unroll. unfold Roll. apply xx.
|
||||
Qed.
|
||||
|
||||
Definition monic D E (f:D -c> E) := forall x y, f x <= f y -> x <= y.
|
||||
|
||||
Lemma Roll_monic: monic Roll.
|
||||
unfold monic.
|
||||
assert (zz:=@ROLL_monic _ FS). fold Roll in zz. fold DInf in zz. unfold VInf.
|
||||
apply zz.
|
||||
Qed.
|
||||
|
||||
Lemma Unroll_monic: monic Unroll.
|
||||
assert (yy:=@UNROLL_monic _ FS). fold Unroll in yy. fold DInf in yy.
|
||||
unfold monic. apply yy.
|
||||
Qed.
|
||||
|
||||
Lemma inj_monic: forall D E (f:D -c> E), monic f -> forall x y, f x == f y -> x == y.
|
||||
intros D E f mm. unfold monic in mm. intros x y.
|
||||
assert (xx:= mm x y). specialize (mm y x).
|
||||
intros [xy1 xy2]. split ; auto.
|
||||
Qed.
|
||||
|
||||
Fixpoint SemEnv n : cpo :=
|
||||
match n with
|
||||
| O => DOne
|
||||
| S n => SemEnv n *c* VInf
|
||||
end.
|
||||
|
||||
Fixpoint projenv (m n:nat) : (m < n) -> SemEnv n -c> VInf :=
|
||||
match m,n with
|
||||
| m,O => fun inconsistant => match (lt_n_O m inconsistant) with end
|
||||
| O, S n => fun _ => SND
|
||||
| S m, S n => fun h => projenv (lt_S_n _ _ h) << FST
|
||||
end.
|
||||
|
||||
(*
|
||||
Lemma K_com2: forall E F (f:E -C-> F),
|
||||
EV << PROD_fun (@K E _ f) ID == f.
|
||||
intros. apply fcont_eq_intro. auto.
|
||||
Qed.
|
||||
|
||||
Lemma uncurry_PROD_fun: forall D D' E F G f g1 (g2:G -C-> D') h,
|
||||
uncurry D E F f << PROD_fun (g1 << g2) h == uncurry D' E F (f << g1) << PROD_fun g2 h.
|
||||
intros. apply fcont_eq_intro. auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism D E F : (@curry D E F)
|
||||
with signature (@Oeq (Dprod D E -C-> F)) ==> (@Oeq (D -C-> E -C-> F))
|
||||
as Curry_eq_compat.
|
||||
intros f0 f1 feq. refine (fcont_eq_intro _). intros d. refine (fcont_eq_intro _). intros e.
|
||||
repeat (rewrite curry_simpl). auto.
|
||||
Qed.
|
||||
*)
|
||||
|
||||
Definition Dlift : (VInf -C-> DInf) -C-> DInf -C-> DInf := curry (Roll << EV <<
|
||||
PROD_map (KLEISLI << (fcont_COMP VInf DInf _ Unroll)) Unroll).
|
||||
|
||||
Definition Dapp := EV << PROD_map (SUM_fun (PBot : Discrete nat -C-> DInf -C-> DL VInf)
|
||||
(fcont_COMP _ _ _ Unroll))
|
||||
(Roll << eta) : Dprod VInf VInf -C-> DL VInf.
|
||||
|
||||
Definition zeroCasem : Discrete nat -m> Dsum DOne (Discrete nat).
|
||||
exists (fun (n:Discrete nat) => match n with | O => @INL DOne _ tt | S m => @INR _ (Discrete nat) m end).
|
||||
unfold monotonic. intros x y xy. assert (x = y) as E by auto.
|
||||
rewrite E in *. auto.
|
||||
Defined.
|
||||
|
||||
Definition zeroCase : Discrete nat -c> Dsum DOne (Discrete nat).
|
||||
exists zeroCasem. unfold continuous. intros c.
|
||||
refine (Ole_trans _ (le_lub _ 0)). auto.
|
||||
Defined.
|
||||
|
||||
Lemma zeroCase_simplS: forall n, zeroCase (S n) = @INR _ (Discrete nat) n.
|
||||
intros n ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma zeroCase_simplO: zeroCase O = @INL DOne _ tt.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Definition DiscProdm S T : Dprod (Discrete S) (Discrete T) -m> Discrete (S * T).
|
||||
intros S T. exists (fun (p:Dprod (Discrete S) (Discrete T)) => (fst p, snd p)).
|
||||
unfold monotonic. intros. destruct x. destruct y. inversion H. simpl in *.
|
||||
rewrite H0. rewrite H1. auto.
|
||||
Defined.
|
||||
|
||||
Definition DiscProd S T : Dprod (Discrete S) (Discrete T) -c> Discrete (S * T).
|
||||
intros S T. exists (DiscProdm S T).
|
||||
unfold continuous. auto.
|
||||
Defined.
|
||||
|
||||
Fixpoint SemVal v n (vt : VTyping n v) : (SemEnv n -C-> VInf) :=
|
||||
match vt with
|
||||
| TNUM n => INL << (@K _ (Discrete nat) n)
|
||||
| TVAR m nthm => projenv nthm
|
||||
| TLAMBDA _ e => INR << Dlift << curry (Roll << SemExp e)
|
||||
end
|
||||
with SemExp pe n (e : ETyping n pe) : (SemEnv n -C-> DL VInf) :=
|
||||
match e with
|
||||
| TAPP _ _ v1 v2 => Dapp << <| SemVal v1, SemVal v2 |>
|
||||
| TVAL _ v => eta << SemVal v
|
||||
| TLET _ _ e1 e2 => EV << <| curry (KLEISLIR (SemExp e2)), SemExp e1 |>
|
||||
| TOP op _ _ v1 v2 =>
|
||||
uncurry (Operator2 (eta << INL << (SimpleOp (fun p => op (fst p) (snd p))) << DiscProd _ _)) <<
|
||||
<| SUM_fun eta (PBot : _ -C-> DL _) << SemVal v1,
|
||||
SUM_fun eta (PBot : _ -C-> DL _) << SemVal v2 |>
|
||||
| TIFZ _ _ _ v e1 e2 => EV <<
|
||||
<| SUM_fun (SUM_fun (K _ (SemExp e1)) (K _ (SemExp e2)) << zeroCase)
|
||||
(PBot : _ -C-> SemEnv n -C-> DL VInf) << SemVal v, ID |>
|
||||
end.
|
||||
|
||||
Lemma Sem_eq:
|
||||
(forall n v (tv0:VTyping n v) n' k x
|
||||
(tv1:VTyping n' (shiftV x k v)) env0 env1,
|
||||
(forall i (nli:i < n) nlik,
|
||||
projenv nli env0 ==
|
||||
@projenv (i+(if bleq_nat x i then k else 0)) _ nlik env1) ->
|
||||
SemVal tv0 env0 == SemVal tv1 env1) /\
|
||||
(forall n e (te0:ETyping n e) n' k x
|
||||
(te1:ETyping n' (shiftE x k e)) env0 env1,
|
||||
(forall i (nli:i < n) nlik,
|
||||
projenv nli env0 == @projenv (i+(if bleq_nat x i then k else 0)) _ nlik env1) ->
|
||||
SemExp te0 env0 == SemExp te1 env1).
|
||||
Proof.
|
||||
apply Typing_ind.
|
||||
intros n m l n' k x.
|
||||
simpl. case_eq (bleq_nat x m) ; intros xm tv1; dependent inversion tv1 ; clear tv1 ; intros env0 env1 C.
|
||||
simpl.
|
||||
specialize (C _ l). rewrite xm in C. specialize (C l0). apply C.
|
||||
specialize (C _ l). rewrite xm in C. assert (m+0 = m) as A by omega. rewrite A in C.
|
||||
apply (C l0).
|
||||
|
||||
intros n m n' k x tv1. dependent inversion tv1 ; clear tv1 ; intros env0 env1 C.
|
||||
simpl.
|
||||
repeat (rewrite fcont_comp_simpl).
|
||||
refine (app_eq_compat (Oeq_refl (@INL (Discrete nat) (DInf -C-> DInf))) _).
|
||||
repeat (rewrite K_simpl). auto.
|
||||
|
||||
intros n e te IHe n' k x tv1 ; dependent inversion tv1 ; clear tv1 ; intros env0 env1 C.
|
||||
simpl. repeat (rewrite fcont_comp_simpl).
|
||||
refine (app_eq_compat (Oeq_refl (@INR (Discrete nat) (DInf -C-> DInf))) _).
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
refine (fcont_eq_intro _).
|
||||
intros v. repeat (rewrite curry_simpl).
|
||||
specialize (IHe _ _ _ e1 (env0,v) (env1,v)).
|
||||
repeat (rewrite fcont_comp_simpl).
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
apply IHe. intros i ; case i ; clear i ; simpl. intros _ _.
|
||||
repeat (rewrite SND_simpl). auto.
|
||||
intros i nli nlik. repeat (rewrite fcont_comp_simpl ; rewrite FST_simpl).
|
||||
simpl. apply C.
|
||||
|
||||
intros n v tv IHv n' k x te ; dependent inversion te ; intros env0 env1 C.
|
||||
simpl. repeat (rewrite fcont_comp_simpl).
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
apply IHv. apply C.
|
||||
|
||||
intros n v1 v2 tv1 IH1 tv2 IH2 n' k x ta ; dependent inversion ta ; intros env0 env1 C.
|
||||
simpl.
|
||||
repeat (rewrite fcont_comp_simpl).
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
repeat (rewrite PROD_fun_simpl). repeat (rewrite PAIR_simpl). simpl.
|
||||
specialize (IH1 _ _ _ v env0 env1 C).
|
||||
specialize (IH2 _ _ _ v0 env0 env1 C).
|
||||
refine (pair_eq_compat IH1 IH2).
|
||||
|
||||
intros n e1 e2 te1 IH1 te2 IH2 n' k x ta ; dependent inversion ta ; intros env0 env1 C.
|
||||
simpl.
|
||||
repeat (rewrite fcont_comp_simpl).
|
||||
repeat (rewrite PROD_fun_simpl).
|
||||
repeat (rewrite curry_simpl).
|
||||
specialize (IH1 _ _ _ e env0 env1 C).
|
||||
refine (Oeq_trans (app_eq_compat (Oeq_refl _) (pair_eq_compat (Oeq_refl _) IH1)) _).
|
||||
refine (KLEISLIR_eq _ _ _).
|
||||
intros v0 v1 sv0 sv1.
|
||||
specialize (IH2 _ _ _ e4 (env0,v0) (env1,v1)).
|
||||
refine (IH2 _). intros i. case i ; clear i.
|
||||
simpl. intros _ _. repeat (rewrite SND_simpl). simpl.
|
||||
apply (vinj (Oeq_trans (Oeq_sym sv0) sv1)).
|
||||
|
||||
simpl. intros i nli nlik.
|
||||
repeat (rewrite fcont_comp_simpl ; rewrite FST_simpl). simpl.
|
||||
apply C.
|
||||
|
||||
intros vv v1. exists vv. auto.
|
||||
intros vv v1. exists vv. auto.
|
||||
|
||||
intros n v e1 e2 tv IHv te1 IH1 te2 IH2 n' k x tif env0 env1 C.
|
||||
dependent inversion tif.
|
||||
simpl. repeat (rewrite fcont_comp_simpl). repeat (rewrite PROD_fun_simpl).
|
||||
repeat (rewrite ID_simpl). simpl.
|
||||
rewrite EV_simpl. rewrite EV_simpl.
|
||||
repeat (rewrite fcont_comp_simpl).
|
||||
specialize (IH1 _ _ _ e env0 env1 C).
|
||||
specialize (IHv _ _ _ v1 env0 env1 C).
|
||||
specialize (IH2 _ _ _ e4 env0 env1 C).
|
||||
refine (Oeq_trans _ (App_eq_compat (app_eq_compat (Oeq_refl (SUM_fun _ _)) IHv) (Oeq_refl env1))).
|
||||
repeat (rewrite SUM_fun_simpl). fold VInf. case (SemVal tv env0) ; intros tt ; simpl ; case tt.
|
||||
unfold INL ; auto. unfold INR ; auto. intros ff _. auto.
|
||||
|
||||
(* TOP *)
|
||||
intros n op v1 v2 tv1 IH1 tv2 IH2 n' k x te1 env0 env1 C.
|
||||
dependent inversion te1. simpl.
|
||||
repeat (rewrite fcont_comp_simpl).
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
rewrite PROD_fun_simpl. rewrite PROD_fun_simpl. simpl.
|
||||
repeat (rewrite fcont_comp_simpl).
|
||||
refine (@pair_eq_compat (DL (Discrete nat)) (DL (Discrete nat)) _ _ _ _ _ _).
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
apply (IH1 _ _ _ v env0 env1 C).
|
||||
refine (app_eq_compat (Oeq_refl _) _).
|
||||
apply (IH2 _ _ _ v4 env0 env1 C).
|
||||
Qed.
|
||||
|
||||
Lemma shift_nil: (forall v k, shiftV k 0 v = v) /\
|
||||
(forall e k, shiftE k 0 e = e).
|
||||
apply ExpValue_ind ; intros ; simpl ;
|
||||
try (try (rewrite (H0 k)) ; try (rewrite (H1 k)) ; try (rewrite (H2 k));
|
||||
try (rewrite (H k)) ; try (rewrite (H (S k))) ; try (rewrite (H0 (S k))) ;
|
||||
try (rewrite (H (S (S k)))) ; auto).
|
||||
case_eq (bleq_nat k n) ; intros beq ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma projenv_irr: forall n i (p1:i < n) (p2:i < n), projenv p1 == projenv p2.
|
||||
intros n. induction n.
|
||||
intros i incon. inversion incon.
|
||||
|
||||
intros i. case i. simpl. auto. clear i.
|
||||
intros i p1 p2. simpl. specialize (IHn _ (lt_S_n i n p1) (lt_S_n i n p2)).
|
||||
rewrite IHn. auto.
|
||||
Qed.
|
||||
|
||||
Lemma Sem_eqV: forall n v (ty0: VTyping n v) (ty1: VTyping n v), SemVal ty0 == SemVal ty1.
|
||||
intros n v tv1 tv2.
|
||||
simpl. refine (fcont_eq_intro _).
|
||||
intros e.
|
||||
assert (xx := proj1 Sem_eq _ _ tv1).
|
||||
specialize (xx n 0 0). rewrite (proj1 shift_nil) in xx.
|
||||
specialize (xx tv2 e e). apply xx.
|
||||
intros i. simpl. assert (i + 0 = i) as A by omega. rewrite A.
|
||||
intros nli nlik.
|
||||
apply (fcont_eq_elim (projenv_irr nli nlik) e).
|
||||
Qed.
|
||||
|
||||
Lemma Sem_eqE: forall n v (ty0: ETyping n v) (ty1: ETyping n v), SemExp ty0 == SemExp ty1.
|
||||
intros n v tv1 tv2.
|
||||
simpl. refine (fcont_eq_intro _).
|
||||
intros e.
|
||||
assert (xx := proj2 Sem_eq _ _ tv1).
|
||||
specialize (xx n 0 0). rewrite (proj2 shift_nil) in xx.
|
||||
specialize (xx tv2 e e). apply xx.
|
||||
intros i. simpl. assert (i + 0 = i) as A by omega. rewrite A.
|
||||
intros nli nlik.
|
||||
apply (fcont_eq_elim (projenv_irr nli nlik) e).
|
||||
Qed.
|
||||
|
514
papers/domains-2009/unisound.v
Executable file
514
papers/domains-2009/unisound.v
Executable file
|
@ -0,0 +1,514 @@
|
|||
Require Import PredomAll.
|
||||
Require Import unisem.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
Fixpoint NSem n n' : (forall i, i < n -> SemEnv n' -C-> VInf) -> SemEnv n' -C-> SemEnv n :=
|
||||
match n as n0 return (forall i, i < n0 -> SemEnv n' -C-> VInf) -> SemEnv n' -C-> SemEnv n0 with
|
||||
| O => (fun _ => @K (SemEnv n') DOne tt)
|
||||
| S n => (fun C => PROD_fun (NSem (fun i nth => C (S i) (lt_n_S _ _ nth))) (C 0 (lt_O_Sn _)))
|
||||
end.
|
||||
|
||||
Lemma NSem_proof_irr: forall n n'
|
||||
(tvs1:forall i, i < n -> SemEnv n' -C-> VInf)
|
||||
(tvs2:forall i, i < n -> SemEnv n' -C-> VInf),
|
||||
(forall i (p:i < n), tvs1 i p == tvs2 i p) ->
|
||||
NSem tvs1 == NSem tvs2.
|
||||
Proof.
|
||||
intros n. induction n ; auto.
|
||||
simpl.
|
||||
intros n' tvs1 tvs2 C. refine (Dprod_unique _ _).
|
||||
repeat (rewrite FST_PROD_fun).
|
||||
apply (IHn n'). intros. apply (C (S i) (lt_n_S _ _ p)).
|
||||
repeat (rewrite SND_PROD_fun).
|
||||
refine (C 0 _).
|
||||
Qed.
|
||||
|
||||
Fixpoint getTyping n n' vs : n = length vs -> (forall i v, nth_error vs i = value v -> VTyping n' v) ->
|
||||
forall i, i < n -> SemEnv n' -C-> VInf :=
|
||||
match n as n0, vs as vs0 return n0 = length vs0 -> (forall i v, nth_error vs0 i = value v -> VTyping n' v) ->
|
||||
forall i, i < n0 -> SemEnv n' -C-> VInf with
|
||||
| O,_ => (fun _ _ i incon => match (lt_n_O i incon) with end)
|
||||
| S n, nil => (fun l _ i _ => match (O_S _ (sym_eq l)) with end)
|
||||
| S n, v :: vs => (fun l C i =>
|
||||
match i with
|
||||
| O => (fun _ => SemVal (C 0 v (refl_equal _)))
|
||||
| S i => (fun li => getTyping (eq_add_S _ _ l) (fun i v nth => C (S i) v nth) (lt_S_n _ _ li))
|
||||
end)
|
||||
end.
|
||||
|
||||
(*
|
||||
Definition getTyping: forall n n' vs (l:n = length vs)
|
||||
(tvs:forall i v, nth_error vs i = value v -> VTyping n' v),
|
||||
(forall i, i < n -> SemEnv n' -C-> VInf).
|
||||
intros n.
|
||||
induction n.
|
||||
intros n' vs l C i p. assert False as F by omega. inversion F.
|
||||
|
||||
intros n' vs. destruct vs.
|
||||
intros incon. simpl in incon. assert False as F by omega. inversion F.
|
||||
intros l C i. case i. clear i. intros _.
|
||||
apply (SemVal (C 0 v (refl_equal _))).
|
||||
|
||||
clear i ; intros i si.
|
||||
specialize (IHn n' vs (eq_add_S _ _ l)).
|
||||
eapply (IHn (fun i v nth => C (S i) v nth) i (lt_S_n _ _ si)).
|
||||
Defined.
|
||||
*)
|
||||
|
||||
Lemma getTyping_irr: forall n n' vs (l1 l2:n = length vs)
|
||||
(tvs1: forall i v, nth_error vs i = value v -> VTyping n' v)
|
||||
(tvs2: forall i v, nth_error vs i = value v -> VTyping n' v)
|
||||
i (nth1 nth2 : i < n), @getTyping n n' vs l1 tvs1 i nth1 == getTyping l2 tvs2 nth2.
|
||||
intros n. induction n. intros n' vs l1 l2 tvs1 tvs2 i incon. assert False as F by omega. inversion F.
|
||||
|
||||
intros n' vs. destruct vs.
|
||||
intros l1. simpl in l1. assert False as F by omega. inversion F.
|
||||
intros l1 l2 tvs1 tvs2 i nth1 nth2. destruct i. simpl. refine (Sem_eqV _ _).
|
||||
|
||||
simpl. refine (IHn _ _ _ _ _ _ _ _ _).
|
||||
Qed.
|
||||
|
||||
Lemma NSem_ext: forall n n' f g, (forall a b, f a b == g a b) -> @NSem n n' f == NSem g.
|
||||
intros n n' f g C.
|
||||
induction n. simpl. auto.
|
||||
simpl. refine (Dprod_unique _ _).
|
||||
repeat (rewrite FST_PROD_fun).
|
||||
apply IHn. apply (fun a ll => C (S a) (lt_n_S _ _ ll)).
|
||||
repeat (rewrite SND_PROD_fun).
|
||||
refine (C 0 _).
|
||||
Qed.
|
||||
|
||||
Lemma projNSem: forall n' v (tv:VTyping n' v) n vs (l:n = length vs) m
|
||||
(nm : m < n)
|
||||
(nthv: nth_error vs m = value v)
|
||||
(tvs : forall (i : nat) (v : Value),
|
||||
nth_error vs i = value v -> VTyping n' v),
|
||||
@projenv m n nm << NSem (getTyping l tvs) == SemVal tv.
|
||||
Proof.
|
||||
intros n' v tv n.
|
||||
induction n. intros vs ntht m. intros incon. assert False as F by omega. inversion F.
|
||||
intros vs l m. destruct m. destruct vs. simpl in l. assert False as F by omega. inversion F.
|
||||
simpl. intros _ vv. inversion vv.
|
||||
generalize l. rewrite H0 in *. clear H0 v0 l. intros l tvs. rewrite SND_PROD_fun.
|
||||
refine (Sem_eqV _ _).
|
||||
|
||||
destruct vs. simpl in l. assert False as F by omega. inversion F.
|
||||
intros nm nth tvs. simpl. rewrite fcont_comp_assoc. rewrite FST_PROD_fun.
|
||||
refine (Oeq_trans _ (IHn vs (eq_add_S _ _ l) m (lt_S_n m n nm) nth (fun i v nth => tvs (S i) v nth))).
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
refine (NSem_ext _). intros. refine (getTyping_irr _ _ _ _ _ _).
|
||||
Qed.
|
||||
|
||||
Lemma NSemShift: forall n n' vs (l1:n = length vs) (l2:n = length (map (shiftV 0 1) vs))
|
||||
(tvs1:forall i v, nth_error vs i = value v -> VTyping n' v)
|
||||
(tvs2:forall i v, nth_error (map (shiftV 0 1) vs) i = value v -> VTyping (S n') v),
|
||||
NSem (getTyping l1 tvs1) << FST == NSem (getTyping l2 tvs2).
|
||||
Proof.
|
||||
intros n. induction n.
|
||||
intros n' vs. destruct vs. intros l1 l2 tvs1 tvs2.
|
||||
simpl. rewrite <- K_com. auto.
|
||||
intros l. simpl in l. assert False as F by omega. inversion F.
|
||||
intros n' vs. destruct vs. intros l. simpl in l. assert False as F by omega. inversion l.
|
||||
intros l1 l2 tvs1 tvs2.
|
||||
simpl. refine (Dprod_unique _ _).
|
||||
rewrite <- fcont_comp_assoc. repeat (rewrite FST_PROD_fun).
|
||||
|
||||
refine (Oeq_trans _ (Oeq_trans (IHn n' vs (eq_add_S _ _ l1) (eq_add_S _ _ l2) (fun i => tvs1 (S i)) (fun i => tvs2 (S i))) _)).
|
||||
refine (fcont_comp_eq_compat _ (Oeq_refl _)).
|
||||
refine (NSem_ext _). intros. apply getTyping_irr.
|
||||
refine (@NSem_ext n (S n') _ _ _). intros. refine (getTyping_irr _ _ _ _ _ _).
|
||||
|
||||
rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun.
|
||||
rewrite SND_PROD_fun.
|
||||
|
||||
refine (fcont_eq_intro _). intros E.
|
||||
rewrite (fcont_comp_simpl).
|
||||
apply (proj1 Sem_eq _ _ (tvs1 0 v (refl_equal (value v))) (S n') 1 0
|
||||
(tvs2 0 (shiftV 0 1 v) (refl_equal (value (shiftV 0 1 v))))
|
||||
(FST E) E).
|
||||
intros i nli.
|
||||
case_eq (bleq_nat 0 i).
|
||||
assert (i+1 = S i) as A by omega. rewrite A. clear A.
|
||||
intros _ nlik.
|
||||
simpl. rewrite fcont_comp_simpl.
|
||||
refine (fcont_eq_elim (projenv_irr _ _) _).
|
||||
|
||||
intros incon. inversion incon.
|
||||
Qed.
|
||||
|
||||
Lemma Substitution:
|
||||
(forall n v (tv : VTyping n v) n' vs (l:n = length vs)
|
||||
(tvs:forall i v, nth_error vs i = value v -> VTyping n' v),
|
||||
SemVal tv << NSem (getTyping l tvs) == SemVal (subst_typingV tv l tvs)) /\
|
||||
(forall n e (te : ETyping n e) n' vs (l:n = length vs)
|
||||
(tvs:(forall i v, nth_error vs i = value v -> VTyping n' v)),
|
||||
SemExp te << NSem (getTyping l tvs) == SemExp (subst_typingE te l tvs)).
|
||||
apply Typing_ind.
|
||||
|
||||
(* TVAR *)
|
||||
intros n m ntht n' vs l tvs.
|
||||
assert (exists v, nth_error vs m = value v).
|
||||
clear tvs n'.
|
||||
rewrite l in ntht. clear n l. generalize vs ntht. clear ntht vs. induction m.
|
||||
intros vs l. destruct vs. simpl in l. assert False as F by omega. inversion F.
|
||||
exists v. auto.
|
||||
intros vs l. destruct vs. simpl in l. assert False as F by omega. inversion F.
|
||||
apply IHm. simpl in l. omega.
|
||||
|
||||
destruct H as [v nthv]. simpl.
|
||||
assert (tv := tvs m v nthv).
|
||||
assert (ss:=Sem_eqV (subst_typingV (TVAR ntht) l tvs)).
|
||||
assert (tv2 : VTyping n' (substVal vs (VAR m))).
|
||||
simpl. rewrite nthv. simpl. apply tv.
|
||||
rewrite (ss tv2). clear ss.
|
||||
generalize tv2. clear tv2. simpl substVal. rewrite nthv. simpl.
|
||||
intros tv2. refine (Oeq_trans _ (Sem_eqV tv tv2)). clear tv2.
|
||||
apply (projNSem tv) ; auto.
|
||||
|
||||
(* TNUM *)
|
||||
intros n m n' vs l tvs. simpl. rewrite fcont_comp_assoc. rewrite <- K_com. auto.
|
||||
|
||||
(* LAMBDA *)
|
||||
intros n body tyb IH n' vs l tvs.
|
||||
|
||||
assert (tt:=subst_typingV (TLAMBDA tyb) l tvs).
|
||||
refine (Oeq_trans _ (Sem_eqV tt _)).
|
||||
generalize tt. clear tt. simpl substVal.
|
||||
simpl.
|
||||
|
||||
specialize (IH (S n') (VAR 0 :: (map (shiftV 0 1) vs))).
|
||||
simpl length in IH. assert (S n = length (VAR 0 :: map (shiftV 0 1) vs)).
|
||||
simpl. rewrite map_length. auto.
|
||||
specialize (IH H).
|
||||
assert (forall i v, nth_error (VAR 0:: map (shiftV 0 1) vs) i = value v -> VTyping (S n') v).
|
||||
intros i v. destruct i. simpl. intros ta. inversion ta.
|
||||
eapply TVAR. omega.
|
||||
simpl. specialize (tvs i). rewrite nth_error_map.
|
||||
case_eq (nth_error vs i). intros vv. intros nthvv. specialize (tvs _ nthvv).
|
||||
intros vvv. inversion vvv.
|
||||
assert (xx:=WeakeningV tvs 1 0). assert (n'+1 = S n') as A by omega. rewrite A in xx.
|
||||
apply xx.
|
||||
intros ; discriminate.
|
||||
specialize (IH H0).
|
||||
|
||||
intros tt.
|
||||
|
||||
dependent inversion tt. clear e H2.
|
||||
simpl. repeat (rewrite fcont_comp_assoc).
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
|
||||
apply curry_unique.
|
||||
|
||||
assert (PROD_map (curry (Roll << SemExp tyb) << NSem (getTyping l tvs)) ID ==
|
||||
PROD_map (curry (Roll << SemExp tyb)) ID << PROD_fun (NSem (getTyping l tvs) << FST) (SND (D2:=VInf))).
|
||||
unfold PROD_map.
|
||||
refine (Dprod_unique _ _).
|
||||
rewrite FST_PROD_fun.
|
||||
repeat (rewrite <- fcont_comp_assoc).
|
||||
rewrite FST_PROD_fun.
|
||||
repeat (rewrite fcont_comp_assoc). rewrite FST_PROD_fun. auto.
|
||||
rewrite SND_PROD_fun. repeat (rewrite <- fcont_comp_assoc). rewrite SND_PROD_fun.
|
||||
repeat (rewrite fcont_comp_assoc). rewrite SND_PROD_fun. auto.
|
||||
rewrite H1. rewrite <- fcont_comp_assoc. rewrite <- curry_com.
|
||||
clear H1. rewrite fcont_comp_assoc.
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
|
||||
rewrite <- (Sem_eqE e0) in IH.
|
||||
rewrite <- IH. clear IH tt e0.
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
|
||||
clear tyb body. simpl.
|
||||
refine (Dprod_unique _ _).
|
||||
repeat (rewrite <- fcont_comp_assoc).
|
||||
repeat (rewrite FST_PROD_fun).
|
||||
|
||||
assert (ss := NSemShift (n':= n') l (eq_add_S _ _ H) tvs).
|
||||
assert (forall (i : nat) (v : Value),
|
||||
nth_error (map (shiftV 0 1) vs) i = value v ->
|
||||
VTyping (S n') v). intros i v.
|
||||
rewrite nth_error_map. case_eq (nth_error vs i). intros vv nth va. inversion va.
|
||||
assert (xx:=WeakeningV (tvs i _ nth) 1 0). assert (S n' = n'+1) as A by omega. rewrite A.
|
||||
apply xx.
|
||||
intros _ incon. inversion incon.
|
||||
|
||||
refine (Oeq_trans _ (Oeq_trans (Oeq_sym (ss H1)) _)). clear ss.
|
||||
refine (@NSem_ext n (S n') _ _ _). intros. refine (getTyping_irr _ _ _ _ _ _).
|
||||
auto.
|
||||
rewrite SND_PROD_fun.
|
||||
rewrite SND_PROD_fun.
|
||||
|
||||
assert (tv:= @TVAR (S n') 0 (lt_O_Sn _)).
|
||||
rewrite <- (@Sem_eqV (S n') _ tv).
|
||||
dependent inversion tv. simpl.
|
||||
auto.
|
||||
|
||||
intros n v tv IH n' vs l tvs.
|
||||
assert (tt:= subst_typingE (TVAL tv) l tvs).
|
||||
rewrite <- (Sem_eqE tt).
|
||||
dependent inversion tt.
|
||||
simpl. rewrite fcont_comp_assoc.
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
rewrite IH. rewrite <- (Sem_eqV v1). auto.
|
||||
|
||||
(* TAPP *)
|
||||
intros n e1 e2 te1 IH1 te2 IH2 E' vs l tvs.
|
||||
assert (tt:= subst_typingE (TAPP te1 te2) l tvs).
|
||||
rewrite <- (Sem_eqE tt).
|
||||
dependent inversion tt.
|
||||
simpl. rewrite fcont_comp_assoc. refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
refine (Dprod_unique _ _).
|
||||
rewrite <- fcont_comp_assoc.
|
||||
repeat (rewrite FST_PROD_fun).
|
||||
rewrite (IH1 _ _ l tvs). rewrite <- (Sem_eqV v) ; auto.
|
||||
rewrite <- fcont_comp_assoc. repeat (rewrite SND_PROD_fun).
|
||||
rewrite (IH2 _ _ l tvs). rewrite <- (Sem_eqV v0) ; auto.
|
||||
|
||||
(* TLET *)
|
||||
intros n e1 e2 te2 IH2 te1 IH1 n' vs l tvs.
|
||||
assert (tt:= subst_typingE (TLET te2 te1) l tvs).
|
||||
rewrite <- (Sem_eqE tt).
|
||||
dependent inversion tt.
|
||||
assert (xx:= subst_typingE te2 l tvs).
|
||||
simpl.
|
||||
rewrite fcont_comp_assoc.
|
||||
specialize (IH2 _ _ l tvs).
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
refine (Dprod_unique _ _).
|
||||
rewrite <- fcont_comp_assoc. repeat (rewrite FST_PROD_fun).
|
||||
refine (curry_unique _).
|
||||
assert (PROD_map (curry (KLEISLIR (SemExp te1)) << NSem (getTyping l tvs)) ID ==
|
||||
PROD_map (curry (KLEISLIR (SemExp te1))) ID << PROD_map (NSem (getTyping l tvs)) (ID (D:=DL VInf))).
|
||||
unfold PROD_map.
|
||||
refine (Dprod_unique _ _).
|
||||
repeat (rewrite FST_PROD_fun).
|
||||
rewrite <- fcont_comp_assoc. rewrite FST_PROD_fun.
|
||||
repeat (rewrite fcont_comp_assoc). rewrite FST_PROD_fun.
|
||||
auto. rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun.
|
||||
repeat (rewrite <- fcont_comp_assoc). rewrite SND_PROD_fun.
|
||||
rewrite fcont_comp_assoc. rewrite SND_PROD_fun.
|
||||
rewrite fcont_comp_unitL. rewrite fcont_comp_unitL. auto.
|
||||
|
||||
rewrite H. clear H. rewrite <- fcont_comp_assoc.
|
||||
rewrite <- curry_com.
|
||||
unfold PROD_map.
|
||||
rewrite KLEISLIR_comp.
|
||||
assert (PROD_map (NSem (getTyping l tvs) << FST) ID ==
|
||||
PROD_map (NSem (getTyping l tvs)) ID << PROD_map (FST (D2:=DL VInf)) (ID( D:=VInf))).
|
||||
rewrite PROD_map_PROD_map. rewrite fcont_comp_unitL. auto.
|
||||
rewrite H. rewrite <- fcont_comp_assoc.
|
||||
rewrite <- KLEISLIR_comp. rewrite fcont_comp_unitL.
|
||||
rewrite (fun D E => Dprod_unique (f := <| @FST D E, SND |>) (g:=ID)).
|
||||
rewrite fcont_comp_unitR.
|
||||
clear H.
|
||||
|
||||
refine (KLEISLIR_eq_compat _).
|
||||
specialize (IH1 (S n') (VAR 0::map (shiftV 0 1) vs)).
|
||||
assert (S n = S (length (map (shiftV 0 1) vs))). rewrite map_length. auto.
|
||||
specialize (IH1 H).
|
||||
assert (forall i v,
|
||||
nth_error (VAR 0 :: map (shiftV 0 1) vs) i = value v -> VTyping (S n') v).
|
||||
intros i s. destruct i. simpl.
|
||||
intros vv. inversion vv.
|
||||
eapply TVAR. omega. simpl.
|
||||
intros nthv. rewrite nth_error_map in nthv. case_eq (nth_error vs i).
|
||||
intros v nthvv. rewrite nthvv in nthv. inversion nthv.
|
||||
assert (S n' = n' + 1) as A by omega. rewrite A.
|
||||
eapply (WeakeningV). apply (tvs i v nthvv). intros nth. rewrite nth in nthv. inversion nthv.
|
||||
|
||||
specialize (IH1 H2).
|
||||
rewrite <- (Sem_eqE e4) in IH1. rewrite <- IH1. clear IH1 H1.
|
||||
simpl. unfold PROD_map.
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) (PROD_fun_eq_compat _ _ )).
|
||||
assert (sh := NSemShift l (eq_add_S _ _ H) tvs (fun i => H2 (S i))).
|
||||
rewrite sh. clear sh. refine (@NSem_ext n (S n') _ _ _).
|
||||
intros. refine (getTyping_irr _ _ _ _ _ _).
|
||||
assert (VTyping (S n') (VAR 0)) as tv by (eapply TVAR ; omega).
|
||||
rewrite <- (Sem_eqV tv).
|
||||
dependent inversion tv. simpl. rewrite fcont_comp_unitL. auto.
|
||||
rewrite fcont_comp_unitR. apply FST_PROD_fun. rewrite fcont_comp_unitR. apply SND_PROD_fun.
|
||||
rewrite <- fcont_comp_assoc. rewrite SND_PROD_fun. rewrite SND_PROD_fun.
|
||||
rewrite IH2. refine (Sem_eqE _ _).
|
||||
|
||||
(* TIFZ *)
|
||||
intros n v e1 e2 tv IHv te1 IH1 te2 IH2 n' vs l tvs.
|
||||
specialize (IHv _ _ l tvs). specialize (IH1 _ _ l tvs). specialize (IH2 _ _ l tvs).
|
||||
assert (tt := subst_typingE (TIFZ tv te1 te2) (vs:=vs) l tvs). rewrite <- (Sem_eqE tt).
|
||||
generalize tt. simpl. clear tt. intros tt.
|
||||
dependent inversion tt. clear H0 H1 H2 v0 e0 e3. simpl.
|
||||
rewrite fcont_comp_assoc.
|
||||
rewrite PROD_fun_compl.
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_assoc _ _ _) (Oeq_refl _)).
|
||||
rewrite <- (Sem_eqV v1) in IHv. rewrite IHv. rewrite fcont_comp_unitL.
|
||||
rewrite <- (PROD_fun_eq_compat (Oeq_refl _) (fcont_comp_unitR (NSem (getTyping l tvs)))).
|
||||
|
||||
Lemma xx D E F G (f:E -C-> F -C-> G) (h:D -C-> F) (g: E -C-> D) :
|
||||
EV << PROD_fun f (h << g) ==
|
||||
EV << PROD_fun (curry (uncurry f << SWAP << PROD_map h ID << SWAP)) g.
|
||||
intros ; apply fcont_eq_intro ; auto.
|
||||
Qed.
|
||||
|
||||
rewrite xx. refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
refine (PROD_fun_eq_compat _ (Oeq_refl _)). apply Oeq_sym.
|
||||
refine (curry_unique _). rewrite fcont_comp_assoc. rewrite fcont_comp_assoc.
|
||||
assert (forall D, SWAP << ((NSem (n:=n) (n':=n') (getTyping (vs:=vs) l tvs) *f* ID) << SWAP) ==
|
||||
@ID D *f* (NSem (n:=n) (n':=n') (getTyping (vs:=vs) l tvs))) as eeq by
|
||||
(intros D ; rewrite <- fcont_comp_assoc ; apply fcont_eq_intro ; auto).
|
||||
rewrite eeq. clear eeq.
|
||||
rewrite <- (Sem_eqE e4) in IH2.
|
||||
rewrite <- (Sem_eqE e) in IH1.
|
||||
apply fcont_eq_intro.
|
||||
intros d. repeat (rewrite fcont_comp_simpl). case d ; clear d. intros ds dn.
|
||||
repeat (rewrite PROD_map_simpl). rewrite ID_simpl. simpl. repeat (rewrite ID_simpl). simpl.
|
||||
repeat (rewrite EV_simpl). rewrite uncurry_simpl. repeat (rewrite fcont_comp_simpl).
|
||||
rewrite SUM_fun_simpl.
|
||||
unfold VInf. rewrite SUM_fun_simpl. fold VInf. simpl.
|
||||
case (SemVal v1 ds). intros t. case t ; auto. unfold INL. simpl.
|
||||
auto using (fcont_eq_elim IH1 dn). clear t. intros t. auto using (fcont_eq_elim IH2 dn).
|
||||
auto.
|
||||
|
||||
(* TOP *)
|
||||
intros n op v1 v2 tv1 IH1 tv2 IH2 n' vs l tvs.
|
||||
specialize (IH1 _ _ l tvs). specialize (IH2 _ _ l tvs).
|
||||
simpl subst_typingE. simpl.
|
||||
refine (Oeq_trans (fcont_comp_assoc _ _ _) _).
|
||||
refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
refine (Dprod_unique _ _). rewrite <- fcont_comp_assoc.
|
||||
repeat (rewrite FST_PROD_fun).
|
||||
rewrite <- IH1. clear IH1.
|
||||
rewrite <- fcont_comp_assoc. auto.
|
||||
rewrite <- fcont_comp_assoc.
|
||||
repeat (rewrite SND_PROD_fun). rewrite <- IH2. clear IH2.
|
||||
rewrite <- fcont_comp_assoc. auto.
|
||||
Qed.
|
||||
|
||||
Lemma Soundness: forall e v, (e =>> v) -> forall (et : ETyping 0 e) (vt : VTyping 0 v), SemExp et == eta << SemVal vt.
|
||||
Proof.
|
||||
intros e v ev. induction ev ; intros te ; dependent inversion te.
|
||||
clear H0 v0 te. intros tv. simpl. refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
apply (Sem_eqV v1 tv).
|
||||
|
||||
clear e2 H0 H1 te. intros tv. simpl. assert (ETyping 1 e1) as te1 by (inversion v0 ; auto).
|
||||
rewrite (PROD_fun_eq_compat (Sem_eqV v0 (TLAMBDA te1)) (Oeq_refl _)).
|
||||
clear e0 v0. unfold Dapp.
|
||||
simpl.
|
||||
assert (forall D E F G H (a: D -c> F) (b : E -c> G) (c : H -c> D) (d : H -c> E), a *f* b << <|c, d|> == <| a << c, b << d |>).
|
||||
intros. unfold PROD_map. refine (Dprod_unique _ _). rewrite <- fcont_comp_assoc. repeat (rewrite FST_PROD_fun).
|
||||
rewrite fcont_comp_assoc. rewrite FST_PROD_fun. auto.
|
||||
rewrite <- fcont_comp_assoc. repeat (rewrite SND_PROD_fun). rewrite fcont_comp_assoc. rewrite SND_PROD_fun. auto.
|
||||
rewrite (fun D E => @fcont_comp_assoc D E (Dprod (DInf -C-> (DL VInf)) DInf) (DL VInf) (EV )).
|
||||
rewrite H. clear H.
|
||||
assert (SUM_fun (D1:=Discrete nat) (D2:=DInf -C-> DInf) (D3:=
|
||||
DInf -C-> DL VInf)
|
||||
(K (Discrete nat) (D2:=DInf -C-> DL VInf)
|
||||
(K DInf (D2:=DL VInf) (DL_bot VInf)))
|
||||
(fcont_COMP DInf DInf (DL VInf) Unroll) <<
|
||||
((@INR (Discrete nat) (DInf -C-> DInf) << Dlift) <<
|
||||
curry (D1:=DOne) (D2:=VInf) (D3:=DInf) (Roll << SemExp te1)) ==
|
||||
(fcont_COMP DInf DInf (DL VInf) Unroll) << Dlift << curry (Roll << SemExp te1)).
|
||||
repeat (rewrite <- fcont_comp_assoc). rewrite Dsum_rcom. auto.
|
||||
rewrite (PROD_fun_eq_compat H (Oeq_refl ((Roll << eta) << SemVal v1))). clear H.
|
||||
unfold Dlift.
|
||||
refine (fcont_eq_intro _).
|
||||
intros tt. rewrite fcont_comp_simpl. rewrite PROD_fun_simpl. simpl. rewrite EV_simpl.
|
||||
repeat (rewrite fcont_comp_simpl). rewrite fcont_COMP_simpl.
|
||||
repeat (rewrite fcont_comp_simpl). rewrite curry_simpl.
|
||||
repeat (rewrite fcont_comp_simpl). rewrite PROD_map_simpl. simpl. rewrite EV_simpl.
|
||||
rewrite UR_eq.
|
||||
rewrite fcont_comp_simpl. rewrite fcont_COMP_simpl. rewrite KLEISLI_simpl.
|
||||
rewrite UR_eq.
|
||||
rewrite eta_val. rewrite kleisli_simpl. rewrite kleisli_Val.
|
||||
rewrite <- fcont_comp_simpl. rewrite <- fcont_comp_simpl.
|
||||
assert (xx := subst_typingE te1).
|
||||
assert (1 = length [v2]) as A by (simpl ; auto). specialize (xx 0 _ A).
|
||||
assert ((forall (i : nat) (v' : Value),
|
||||
nth_error [v2] i = Some v' -> VTyping 0 v')) as C. intros i. destruct i. intros vv. simpl. intro incon. inversion incon.
|
||||
rewrite <- H0. apply v1. simpl. intros vv. destruct i ; simpl ; intros incon ; inversion incon.
|
||||
specialize (xx C).
|
||||
specialize (IHev xx).
|
||||
specialize (IHev tv).
|
||||
rewrite <- (fcont_eq_elim IHev tt).
|
||||
refine (app_eq_compat _ (Oeq_refl tt)).
|
||||
assert (yy:=Oeq_trans (proj2 Substitution 1 _ te1 0 _ A C) (Sem_eqE _ xx)).
|
||||
rewrite <- yy.
|
||||
simpl. assert (zz:=Sem_eqV (C 0 v2 (refl_equal (value v2))) v1).
|
||||
rewrite (PROD_fun_eq_compat (Oeq_refl (K DOne (D2:=DOne) Datatypes.tt)) zz).
|
||||
clear zz yy xx IHev. refine (fcont_eq_intro _). intros t. repeat (rewrite fcont_comp_simpl).
|
||||
rewrite curry_simpl. rewrite fcont_comp_simpl. rewrite UR_eq.
|
||||
refine (app_eq_compat (Oeq_refl _) _). rewrite PROD_fun_simpl. simpl. rewrite K_simpl. case tt. auto.
|
||||
|
||||
(* TLET *)
|
||||
intros tv2. simpl.
|
||||
assert (<|curry (KLEISLIR (SemExp e4)), SemExp e |> == (curry (KLEISLIR (SemExp e4)) *f* ID) << <| ID, SemExp e|>).
|
||||
unfold PROD_map.
|
||||
refine (Dprod_unique _ _). rewrite <- fcont_comp_assoc. repeat (rewrite FST_PROD_fun). rewrite fcont_comp_assoc. rewrite FST_PROD_fun.
|
||||
rewrite fcont_comp_unitR. auto.
|
||||
rewrite <- fcont_comp_assoc. repeat (rewrite SND_PROD_fun). rewrite fcont_comp_assoc. rewrite SND_PROD_fun. rewrite fcont_comp_unitL.
|
||||
auto.
|
||||
rewrite H. clear H. rewrite <- fcont_comp_assoc. rewrite <- curry_com.
|
||||
specialize (IHev1 e). assert (tv1 := eval_preserve_type ev1 e).
|
||||
specialize (IHev1 tv1). rewrite (PROD_fun_eq_compat (Oeq_refl ID) IHev1).
|
||||
rewrite KLEISLIR_unit.
|
||||
clear H0 H1 e0 e3 te.
|
||||
assert (1 = length [v1]) as l by (simpl ; auto).
|
||||
assert ((forall (i : nat) (v' : Value),
|
||||
nth_error [v1] i = Some v' -> VTyping 0 v')) as C.
|
||||
intros i ; destruct i. simpl. intros vv incon ; inversion incon. rewrite <- H0. clear H0 incon vv.
|
||||
apply tv1. simpl. intros vv incon ; destruct i ; simpl in incon ; inversion incon.
|
||||
assert (xx:=subst_typingE e4 l C).
|
||||
specialize (IHev2 xx tv2). rewrite <- IHev2.
|
||||
assert (yy:=proj2 Substitution 1 _ e4 0 _ l C).
|
||||
rewrite <- (Sem_eqE xx) in yy. rewrite <- yy.
|
||||
simpl. refine (fcont_comp_eq_compat (Oeq_refl _) _).
|
||||
refine (Dprod_unique _ _).
|
||||
repeat (rewrite FST_PROD_fun).
|
||||
refine (DOne_final _ _).
|
||||
repeat (rewrite SND_PROD_fun). refine (Sem_eqV tv1 _).
|
||||
|
||||
(* TIFZ zero *)
|
||||
intros tv1. clear e0 e3 H0 H1 H2 te. rewrite <- (IHev e tv1). clear IHev.
|
||||
simpl. rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Oeq_refl _) (Sem_eqV v0 (TNUM 0 0))) (Oeq_refl _)).
|
||||
clear v0.
|
||||
simpl.
|
||||
rewrite <- (PROD_fun_eq_compat (fcont_comp_assoc _ _ _) (Oeq_refl _)).
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Dsum_lcom _ _) (Oeq_refl _)) (Oeq_refl _)).
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_assoc _ _ _) (Oeq_refl _)).
|
||||
assert ((zeroCase << K DOne (D2:=Discrete nat) 0) == INL << ID) as zeq by
|
||||
(apply fcont_eq_intro ;intros x ; case x ; auto).
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Oeq_refl _) zeq) (Oeq_refl _)).
|
||||
rewrite <- (PROD_fun_eq_compat (fcont_comp_assoc _ _ _) (Oeq_refl _)).
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Dsum_lcom _ _) (Oeq_refl _)) (Oeq_refl _)).
|
||||
apply fcont_eq_intro ; auto.
|
||||
|
||||
(* TIFZ > 0 *)
|
||||
intros tv2. clear e0 e3 H0 H1 H2 te. rewrite <- (IHev e4 tv2). clear IHev.
|
||||
simpl.
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Oeq_refl _) (Sem_eqV v0 (TNUM 0 _))) (Oeq_refl _)).
|
||||
clear v0.
|
||||
simpl.
|
||||
rewrite <- (PROD_fun_eq_compat (fcont_comp_assoc _ _ _) (Oeq_refl _)).
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Dsum_lcom _ _) (Oeq_refl _)) (Oeq_refl _)).
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_assoc _ _ _) (Oeq_refl _)).
|
||||
assert ((zeroCase << K DOne (D2:=Discrete nat) (S n)) == INR << (K DOne (n:Discrete nat))) as zeq by
|
||||
(apply fcont_eq_intro ; auto).
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Oeq_refl _) zeq) (Oeq_refl _)).
|
||||
rewrite <- (PROD_fun_eq_compat (fcont_comp_assoc _ _ _) (Oeq_refl _)).
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Dsum_rcom _ _) (Oeq_refl _)) (Oeq_refl _)).
|
||||
apply fcont_eq_intro ; auto.
|
||||
|
||||
intros tv. clear op0 H0 H1 H2.
|
||||
dependent inversion tv. clear m H0. simpl.
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Oeq_refl _) (Sem_eqV v (TNUM 0 n1)))
|
||||
(fcont_comp_eq_compat (Oeq_refl _) (Sem_eqV v0 (TNUM 0 n2)))).
|
||||
simpl. clear v v1 v2 te.
|
||||
rewrite <- (PROD_fun_eq_compat (fcont_comp_assoc _ _ _) (fcont_comp_assoc _ _ _)).
|
||||
rewrite (PROD_fun_eq_compat (fcont_comp_eq_compat (Dsum_lcom _ _) (Oeq_refl _)) (fcont_comp_eq_compat (Dsum_lcom _ _) (Oeq_refl _))).
|
||||
refine (fcont_eq_intro _). intros tt.
|
||||
repeat (rewrite fcont_comp_simpl).
|
||||
rewrite PROD_fun_simpl. simpl. repeat (rewrite fcont_comp_simpl).
|
||||
rewrite eta_val. rewrite eta_val. rewrite uncurry_simpl.
|
||||
repeat (rewrite K_simpl). rewrite Operator2_simpl. auto.
|
||||
Qed.
|
||||
|
401
papers/domains-2009/utility.v
Executable file
401
papers/domains-2009/utility.v
Executable file
|
@ -0,0 +1,401 @@
|
|||
(* utility.v
|
||||
|
||||
Generally useful stuff that doesn't belong anywhere else for
|
||||
|
||||
Abstracting Allocation : The New New Thing
|
||||
Nick Benton
|
||||
|
||||
(c) 2006 Microsoft Research
|
||||
|
||||
This also exports a bunch of pervasive stuff from the standard library
|
||||
*)
|
||||
|
||||
Require Export Arith.
|
||||
Require Export EqNat.
|
||||
Require Export Bool.
|
||||
Require Export BoolEq.
|
||||
Require Export Setoid.
|
||||
Require Export Omega.
|
||||
Require Export Bool_nat.
|
||||
Require Export List.
|
||||
Require Export Max.
|
||||
Require Export Min.
|
||||
|
||||
Lemma generalise_recursion : forall (P:nat->Prop), (forall k , P k) <-> (forall k k0, k0<=k -> P k0).
|
||||
Proof.
|
||||
intuition; apply (H k k); trivial.
|
||||
Qed.
|
||||
|
||||
Ltac dcase x := generalize (refl_equal x); pattern x at -1; case x.
|
||||
|
||||
Notation "[ x ; .. ; y ]" := (cons x .. (cons y nil) ..).
|
||||
|
||||
Tactic Notation "hreplace" constr(h) "with" constr(t) := let id:=fresh in (assert (id := t); clear h; rename id into h).
|
||||
|
||||
Tactic Notation "omega_replace" constr(m) "with" constr(n) := replace m with n ; try omega.
|
||||
|
||||
|
||||
(*Tactic Notation "deconstruct" hyp(h) := _deconstruct h.*)
|
||||
Tactic Notation "deconstruct" hyp(h) "as" simple_intropattern(ip) := generalize h; clear h; intros ip.
|
||||
|
||||
Tactic Notation "witness" ident(p) "in" hyp(h) := elim h; clear h; intros p h.
|
||||
|
||||
Axiom Extensionality : forall (A B : Type) (f g : A->B), (forall a:A, f a = g a) <-> f =g.
|
||||
|
||||
Lemma strictPosSucc: forall n, n>0 -> exists m , S m =n.
|
||||
Proof.
|
||||
intros.
|
||||
induction n.
|
||||
destruct H.
|
||||
exists n; trivial.
|
||||
exists m; trivial.
|
||||
exists n; trivial.
|
||||
Qed.
|
||||
|
||||
Ltac unfold_plus:=
|
||||
try repeat rewrite <- plus_n_Sm;
|
||||
try rewrite plus_0_r;
|
||||
try rewrite plus_0_l.
|
||||
|
||||
Tactic Notation (at level 0) "unfold_plus" "in" constr(H) := try repeat rewrite <- plus_n_Sm in H;
|
||||
try rewrite plus_0_r in H;
|
||||
try rewrite plus_0_l in H.
|
||||
|
||||
Lemma beq_nat_neq :forall m n, m<>n -> false = (beq_nat m n).
|
||||
double induction m n; simpl in |-*.
|
||||
intros.
|
||||
tauto.
|
||||
firstorder.
|
||||
firstorder.
|
||||
intros.
|
||||
set(Hn := (H0 n)).
|
||||
assert (n1 <> n).
|
||||
omega.
|
||||
apply (Hn H2).
|
||||
Qed.
|
||||
|
||||
(* bleq_nat, ble_nat ... *)
|
||||
|
||||
|
||||
Fixpoint bleq_nat (n m: nat) {struct n} : bool :=
|
||||
match n with
|
||||
| 0 => true
|
||||
| S n1 => match m with
|
||||
| 0 => false
|
||||
| S m1 => bleq_nat n1 m1
|
||||
end
|
||||
end.
|
||||
|
||||
|
||||
Fixpoint ble_nat (n m : nat) {struct n} : bool :=
|
||||
match n with
|
||||
| 0 => match m with
|
||||
| 0 => false
|
||||
| S _ => true
|
||||
end
|
||||
| S n1 => match m with
|
||||
| 0 => false
|
||||
| S m1 => ble_nat n1 m1
|
||||
end
|
||||
end.
|
||||
|
||||
|
||||
Lemma ble_nat_false : forall m n, m >=n -> false = ble_nat m n.
|
||||
double induction m n.
|
||||
firstorder.
|
||||
firstorder.
|
||||
firstorder.
|
||||
firstorder.
|
||||
Qed.
|
||||
|
||||
Lemma bleq_nat_false : forall m n, m > n -> false = bleq_nat m n.
|
||||
double induction m n.
|
||||
simpl.
|
||||
intro; absurd (0>0).
|
||||
omega.
|
||||
assumption.
|
||||
|
||||
intros.
|
||||
simpl.
|
||||
absurd (0> S n0).
|
||||
omega.
|
||||
assumption.
|
||||
|
||||
intros.
|
||||
simpl.
|
||||
tauto.
|
||||
|
||||
intros.
|
||||
simpl.
|
||||
apply H0.
|
||||
omega.
|
||||
Qed.
|
||||
|
||||
Lemma bleq_nat_true : forall m n, m <= n -> true = bleq_nat m n.
|
||||
double induction m n; intros; simpl.
|
||||
tauto.
|
||||
tauto.
|
||||
absurd (S n <=0).
|
||||
omega.
|
||||
assumption.
|
||||
apply H0.
|
||||
omega.
|
||||
Qed.
|
||||
|
||||
Lemma true_bleq_nat : forall m n, true = bleq_nat m n -> m <= n.
|
||||
double induction m n.
|
||||
simpl.
|
||||
omega.
|
||||
simpl; intros.
|
||||
omega.
|
||||
simpl; intros.
|
||||
discriminate H0.
|
||||
simpl; intros.
|
||||
intuition.
|
||||
Qed.
|
||||
|
||||
Lemma false_bleq_nat : forall m n, false = bleq_nat m n -> m > n.
|
||||
intros.
|
||||
assert (~ m <= n).
|
||||
intro Habs; generalize (bleq_nat_true m n Habs); intro.
|
||||
rewrite <- H0 in H; discriminate H.
|
||||
omega.
|
||||
Qed.
|
||||
|
||||
Lemma ble_nat_true : forall m n, m < n -> true = ble_nat m n.
|
||||
double induction m n.
|
||||
intro; absurd (0<0).
|
||||
omega.
|
||||
assumption.
|
||||
|
||||
intros.
|
||||
simpl.
|
||||
tauto.
|
||||
|
||||
intros.
|
||||
simpl.
|
||||
absurd (S n < 0).
|
||||
omega.
|
||||
assumption.
|
||||
|
||||
intros.
|
||||
simpl.
|
||||
apply H0.
|
||||
omega.
|
||||
Qed.
|
||||
|
||||
Lemma booleq : forall (a b : bool), ((true = a) <-> (true = b)) -> (a = b).
|
||||
intros; dcase a.
|
||||
intuition.
|
||||
intuition.
|
||||
dcase b.
|
||||
rewrite H0 in H2.
|
||||
intro; symmetry; apply H2; symmetry; assumption.
|
||||
tauto.
|
||||
Qed.
|
||||
|
||||
|
||||
Definition compose (A B C : Type) (g : B -> C) (f : A -> B) a := g(f a).
|
||||
|
||||
Implicit Arguments compose [A B C].
|
||||
|
||||
Definition opt_map : forall (A B : Type), (A -> B) -> (option A -> option B) :=
|
||||
fun A B f o =>
|
||||
match o with
|
||||
| None => None
|
||||
| Some a => Some (f a)
|
||||
end.
|
||||
|
||||
Implicit Arguments opt_map [A B].
|
||||
|
||||
Definition opt_lift : forall (A B : Type), (A -> option B) -> option A -> option B :=
|
||||
fun A B f o =>
|
||||
match o with
|
||||
| Some a => f a
|
||||
| None => None
|
||||
end.
|
||||
|
||||
Implicit Arguments opt_lift [A B].
|
||||
|
||||
Section Lists_extras.
|
||||
|
||||
Variable A : Type.
|
||||
|
||||
Set Implicit Arguments.
|
||||
|
||||
(* cons_nth *)
|
||||
|
||||
Fixpoint cons_nth (a:A) (l:list A) (m:nat) {struct l} : list A :=
|
||||
match l with
|
||||
| nil => a :: nil
|
||||
| b :: bs =>
|
||||
match m with
|
||||
| 0 => a :: l
|
||||
| S m' => b :: cons_nth a bs m'
|
||||
end
|
||||
end.
|
||||
|
||||
Lemma cons_nth_S : forall a b l m, cons_nth a (b::l) (S m) = b :: (cons_nth a l m).
|
||||
Proof.
|
||||
firstorder.
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_nil : forall n, nth_error (nil : list A) n = error.
|
||||
Proof.
|
||||
destruct n; trivial.
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_cons : forall (l:list A) n a b, (nth_error l n = Some a) = (nth_error (b::l) (S n) = Some a).
|
||||
simpl; auto.
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_cons_nth_lt : forall (l:list A) n a b m (h:n<m),
|
||||
(nth_error l n = Some a) -> (nth_error (cons_nth b l m) n = Some a).
|
||||
Proof.
|
||||
induction l.
|
||||
intros; rewrite (nth_error_nil n) in H; inversion H.
|
||||
intros until m; case m.
|
||||
intro h; absurd (n < 0); [apply (lt_n_O n) | assumption].
|
||||
case n ;[ trivial | simpl; intros ; apply (IHl n0 a0 b n1) ; [omega |trivial]].
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_cons_nth_ltI : forall (l:list A) n a b m (h : n < m) (hn:n < length l),
|
||||
(nth_error (cons_nth b l m) n = Some a) -> (nth_error l n = Some a).
|
||||
Proof.
|
||||
induction l.
|
||||
intros. simpl in hn. destruct (lt_n_O n). auto.
|
||||
intros. destruct n. simpl in H. destruct m. destruct (lt_irrefl 0) ; auto.
|
||||
auto. simpl. simpl in H ; destruct m. destruct (lt_n_O (S n)). auto.
|
||||
assert (S n < S (length l)). auto.
|
||||
apply (IHl n a0 b m (lt_S_n _ _ h) (lt_S_n _ _ H0)). auto.
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_cons_nth_ge : forall (l:list A) n a b m (h:n>=m),
|
||||
(nth_error l n = Some a) -> (nth_error (cons_nth b l m) (S n) = Some a).
|
||||
Proof.
|
||||
induction l.
|
||||
intro n; rewrite (nth_error_nil n); intros; inversion H.
|
||||
intros until m;case m.
|
||||
trivial.
|
||||
case n.
|
||||
intros n0 h; absurd (0 >= S n0) ; [omega | trivial].
|
||||
intros n0 n1 h; simpl; simpl in IHl.
|
||||
apply IHl; omega.
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_cons_nth_geI : forall (l:list A) n a b m (h:n>=m),
|
||||
(nth_error (cons_nth b l m) (S n) = Some a) -> (nth_error l n = Some a).
|
||||
Proof.
|
||||
induction l.
|
||||
intro n. auto.
|
||||
intros. destruct m. simpl in H. auto.
|
||||
simpl in H. destruct n. destruct (le_not_gt _ _ h).
|
||||
apply neq_O_lt. auto.
|
||||
apply (IHl n a0 b _ (le_S_n _ _ h) H).
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_cons_nth_eq : forall a l n,
|
||||
n <= length l -> nth_error (cons_nth a l n) n = value a.
|
||||
Proof.
|
||||
intros.
|
||||
generalize dependent l.
|
||||
induction n. intros. destruct l ; simpl ; reflexivity.
|
||||
intros.
|
||||
destruct l. simpl in H.
|
||||
destruct (absurd False H (le_Sn_O n)).
|
||||
simpl.
|
||||
assert (length (a0 :: l) = S (length l)). auto.
|
||||
rewrite -> H0 in H.
|
||||
apply (IHn _ (le_S_n _ _ H)).
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_length : forall l n (a:A), nth_error l n = value a -> n < length l.
|
||||
Proof.
|
||||
intros. generalize dependent l.
|
||||
induction n. destruct l. simpl.
|
||||
intros. inversion H.
|
||||
simpl. intros. apply gt_Sn_O.
|
||||
intro l. destruct l. simpl. intro. inversion H.
|
||||
simpl. intro. apply gt_n_S. apply (IHn _ H).
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_lengthI : forall (l:list A) n, n < length l -> exists a, nth_error l n = value a.
|
||||
Proof.
|
||||
intros l. induction l. simpl. intros n incon. inversion incon.
|
||||
intros n ; case n ; clear n. simpl. intros _. exists a. auto.
|
||||
intros n. simpl. intros C ; apply IHl. omega.
|
||||
Qed.
|
||||
|
||||
Lemma cons_nth_length_S: forall a l n, length (cons_nth a l n) = S (length l).
|
||||
Proof.
|
||||
intros. generalize dependent l. induction n.
|
||||
intros. destruct l ; auto.
|
||||
intros. destruct l. auto.
|
||||
simpl. apply eq_S. apply IHn.
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_0 : forall (l:list A) (a:A), nth_error l 0 = Some a <-> (exists l', l = a::l').
|
||||
Proof.
|
||||
intros; split.
|
||||
induction l.
|
||||
simpl; intro.
|
||||
inversion H.
|
||||
simpl; intro.
|
||||
inversion H.
|
||||
exists l;trivial.
|
||||
intro; destruct H.
|
||||
rewrite H.
|
||||
simpl;trivial.
|
||||
Qed.
|
||||
|
||||
Lemma nth_error_Sn : forall (l:list A) (a:A) (n:nat),
|
||||
nth_error l (S n) = Some a <-> (exists l', exists a0, l = a0::l' /\ nth_error l' n = Some a).
|
||||
Proof.
|
||||
intros; split.
|
||||
induction l.
|
||||
simpl; intro.
|
||||
inversion H.
|
||||
simpl; intro.
|
||||
exists l;exists a0;intuition.
|
||||
intro; destruct H as [l' [a0 (H0,H1)]].
|
||||
rewrite H0.
|
||||
simpl;trivial.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma nth_error_map: forall B (f:A -> B) (l:list A) i,
|
||||
nth_error (map f l) i = (match nth_error l i with Some e => Some (f e) | error => error end).
|
||||
Proof.
|
||||
intros B f l.
|
||||
induction l.
|
||||
simpl. intros i. destruct i ; simpl ; auto.
|
||||
intros i. destruct i. simpl. auto.
|
||||
simpl. apply IHl.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma split_list : forall (a:A) l , a :: l = (a :: nil) ++ l.
|
||||
Proof.
|
||||
trivial.
|
||||
Qed.
|
||||
|
||||
Fixpoint takedrop n (l1 l2 : list A) {struct n} := match n with
|
||||
| 0 => (l1,l2)
|
||||
| S n => match l2 with
|
||||
| nil => (l1,l2)
|
||||
| h::t => takedrop n (l1++(h::nil)) t
|
||||
end
|
||||
end.
|
||||
|
||||
Definition fuse (couple:list A * list A) := match couple with (l1,l2) => l1++l2 end.
|
||||
|
||||
Lemma fuse_unfold : forall l1 l2, fuse (l1,l2) = l1++l2.
|
||||
Proof.
|
||||
unfold fuse; intuition.
|
||||
Qed.
|
||||
|
||||
Opaque fuse.
|
||||
|
||||
Unset Implicit Arguments.
|
||||
|
||||
End Lists_extras.
|
1136
papers/domains-2012/Categories.v
Normal file
1136
papers/domains-2012/Categories.v
Normal file
File diff suppressed because it is too large
Load diff
61
papers/domains-2012/Fin.v
Normal file
61
papers/domains-2012/Fin.v
Normal file
|
@ -0,0 +1,61 @@
|
|||
(**********************************************************************************
|
||||
* Fin.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Finite nats and maps *)
|
||||
|
||||
Require Export ssreflect ssrnat.
|
||||
Require Export NaryFunctions.
|
||||
Require Import Program.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
Inductive Fin : nat -> Type :=
|
||||
| FinZ : forall n, Fin n.+1
|
||||
| FinS : forall n, Fin n -> Fin n.+1.
|
||||
|
||||
Definition FMap n D := Fin n -> D.
|
||||
|
||||
(* Head, tail and cons *)
|
||||
Definition tl n D (m:FMap n.+1 D) : FMap n D := fun v => m (FinS v).
|
||||
Definition hd n D (m:FMap n.+1 D) : D := m (FinZ _).
|
||||
|
||||
Definition cons n D (hd:D) (tl : FMap n D) : FMap n.+1 D :=
|
||||
(fun var =>
|
||||
match var in Fin n return FMap n.-1 D -> D with
|
||||
| FinZ _ => fun _ => hd
|
||||
| FinS _ var' => fun tl' => tl' var'
|
||||
end tl).
|
||||
|
||||
Fixpoint nth n D (v : Fin n) : FMap n D -> D :=
|
||||
match v with
|
||||
| FinZ _ => fun m => hd m
|
||||
| FinS _ i => fun m => nth i (tl m)
|
||||
end.
|
||||
|
||||
Fixpoint asTuple n D : FMap n D -> D ^ n :=
|
||||
match n with
|
||||
| 0 => fun m => tt
|
||||
| n.+1 => fun m => (hd m, asTuple (tl m))
|
||||
end.
|
||||
|
||||
Lemma hdCons : forall n D (v : D) (m : FMap n D), hd (cons v m) = v.
|
||||
Proof. by []. Qed.
|
||||
|
||||
Axiom Extensionality : forall n D (m1 m2 : FMap n D), (forall i, m1 i = m2 i) -> m1 = m2.
|
||||
|
||||
Lemma tlCons : forall n D (v : D) (s : FMap n D), tl (cons v s) = s.
|
||||
Proof. move => n D v s. by apply Extensionality. Qed.
|
||||
|
||||
Ltac extFMap i := (apply Extensionality; intros i; dependent destruction i).
|
||||
|
||||
Lemma consEta : forall n D (m:FMap n.+1 D), m = cons (hd m) (tl m).
|
||||
Proof. move => n D m. extFMap i; by []. Qed.
|
||||
|
||||
|
||||
|
947
papers/domains-2012/Finmap.v
Normal file
947
papers/domains-2012/Finmap.v
Normal file
|
@ -0,0 +1,947 @@
|
|||
(**********************************************************************************
|
||||
* Finmap.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Finite maps with comparison on the domain *)
|
||||
|
||||
(*
|
||||
Summary of things:
|
||||
|
||||
Types:
|
||||
compType is is the type of computable linear orders.
|
||||
[compType of nat] is the linear order of natural numbers.
|
||||
|
||||
FinDom (C:compType) (T:Type) : Type
|
||||
is the type finite maps from C to T.
|
||||
If T is an equality type then findom_eqType C T is an equality type.
|
||||
|
||||
fdEmpty : FinDom C T
|
||||
the every undefined finite map.
|
||||
|
||||
dom_fdEmpty : dom fdEmpty = [::]
|
||||
|
||||
findom_undef (f:FinDom C T) (c \notin dom f) : f c = None
|
||||
|
||||
findom_indom f c : (c \in dom f) = isSome (f c)
|
||||
|
||||
updMap (c:T) (t:T) (f:FinDom C T) : FinDom C T
|
||||
add the pair (c,t) to the finite map f. If c \in dom f then f c is
|
||||
changed to t.
|
||||
|
||||
updMap_simpl : updMap c t f c = Seom t
|
||||
|
||||
updMap_simpl2 : c != c' -> updMap c t f c' = f c'
|
||||
|
||||
indomUpdMap c (f:FinDom C T) c' t : c \in dom (updMap c' t f) = c == c' || c \in dom f
|
||||
|
||||
remMap c (f:FinDom C T) : FinDom C T
|
||||
removes c from the domain of f.
|
||||
|
||||
remMap_simpl c (f:FinDom C T) : remMap c f c = None
|
||||
|
||||
remMap_simpl2 c c' (f:FinDom C T) : c != c' -> remMap c f c' = f c'
|
||||
|
||||
indomRemMap c f t : c \in remMap t f = a != t && a \in dom f
|
||||
|
||||
updMapCom f c t c' t' : c != c' -> updMap c t (updMap c' t' f) = updMap c' t' (updMap c t f)
|
||||
|
||||
remUpdMap (f:FinDom C T) c c' t : c != c' -> remMap c (updMap c' t f) = updMap c' t (remMap c f)
|
||||
|
||||
create_findom (C:compType) (T1:Type) (l:seq T0) (f:T0 -> option T1) : uniq l -> sorted l -> FinDom C T1
|
||||
takes a sorted list to a finite map.
|
||||
|
||||
dom_create_findom gives you the domain of a create_findom.
|
||||
|
||||
findom_map (m:FinDom T T') (f:forall x, x \in dom m -> T'') : FinDom T T''
|
||||
maps f over the domain of m.
|
||||
|
||||
dom_findom_map ...
|
||||
|
||||
post_comp (g:T -> option T') (f:FinDom C T) : FinDom C T'
|
||||
post composition of f with g.
|
||||
|
||||
post_comp_simpl g (f:FinDom C T) c : post_comp g f c = option_bind g (f c)
|
||||
|
||||
dom_post_comp g (f:FinDom C T) : dom (post_copm g f) = filter (fun x => isSome (option_bind g (f x))) (dom f)
|
||||
|
||||
*)
|
||||
|
||||
Require Export ssreflect ssrnat ssrbool seq eqtype.
|
||||
Require Import ssrfun.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
Module Comparison.
|
||||
|
||||
Definition axiomAS (T:Type) l := forall x y : T, reflect (x = y) (l x y && l y x).
|
||||
Definition axiomCL T (l:T -> T -> bool) (comp:T -> T -> unit + bool) :=
|
||||
forall x y, ((comp x y == inl _ tt) = l x y && l y x) /\ (comp x y == inr _ true) = l x y && negb (l y x) /\
|
||||
(comp x y == inr _ false) = l y x && negb (l x y).
|
||||
Definition axiomT T l := forall x y z : T, l x y && l y z ==> l x z.
|
||||
|
||||
Record mixin_of (T:Type) : Type :=
|
||||
Mixin { leq : T -> T -> bool;
|
||||
comp : T -> T -> unit + bool;
|
||||
_ : axiomAS leq;
|
||||
_ : axiomT leq;
|
||||
_ : axiomCL leq comp
|
||||
}.
|
||||
|
||||
Notation class_of := mixin_of (only parsing).
|
||||
Section ClassDef.
|
||||
Structure type : Type := Pack {sort : Type; _ : class_of sort; _ : Type}.
|
||||
Local Coercion sort : type >-> Sortclass.
|
||||
Definition class cT := let: Pack _ c _ := cT return class_of cT in c.
|
||||
|
||||
Definition eqop (T:type) (x y : T) := comp (class T) x y == inl _ tt.
|
||||
|
||||
Lemma eqP T : Equality.axiom (@eqop T).
|
||||
case:T. move => T. case => l c A0 A1 A2 T'. unfold eqop. simpl.
|
||||
move => x y. simpl. specialize (A0 x y). have A:=proj1 (A2 x y). simpl in A. rewrite <- A in A0. by apply A0.
|
||||
Qed.
|
||||
|
||||
Lemma leq_refl (T:type) (x:T) : leq (class T) x x.
|
||||
case:T x. move => T. case. move => l c AS Tr CL X. simpl. move => x. case_eq (l x x) => L; first by [].
|
||||
specialize (AS x x). rewrite L in AS. simpl in AS. by case AS.
|
||||
Qed.
|
||||
|
||||
Lemma leq_trans (T:type) (x y z:T) : leq (class T) x y -> leq (class T) y z -> leq (class T) x z.
|
||||
case:T x y z => T. case. move => l c AS Tr CL X. simpl. move => x y z A B.
|
||||
specialize (Tr x y z). rewrite -> A, -> B in Tr. rewrite implyTb in Tr. by apply Tr.
|
||||
Qed.
|
||||
|
||||
Fixpoint sorted (T:type) (s:seq T) : bool :=
|
||||
match s with
|
||||
| nil => true
|
||||
| x::s' => match s' with | nil => true | y::_ => leq (class T) x y && sorted s' end
|
||||
end.
|
||||
|
||||
Lemma leq_seq_trans T x t s : leq (class T) x t -> all (leq (class T) t) s -> all (leq (class T) x) s.
|
||||
elim:s ; first by [].
|
||||
move => e s IH X. specialize (IH X). simpl. move => Y. rewrite (IH (proj2 (andP Y))). rewrite andbT.
|
||||
apply (@leq_trans _ x t e). by rewrite X. by rewrite (proj1 (andP Y)).
|
||||
Qed.
|
||||
|
||||
Lemma ltn_trans T x y t : leq (class T) x y -> leq (class T) t x -> ~~ eqop t x -> ~~ eqop t y.
|
||||
move: T (@eqP T) x y t. case => T. simpl. case => l c AS Tr CL X E. simpl. unfold eqop in E. simpl in E.
|
||||
move => x y t L L'. unfold eqop. simpl. case_eq (c t y) ; last by case.
|
||||
case. move => e. have e':c t y == inl bool tt by rewrite e.
|
||||
have ee:=E _ _ e'. rewrite <- ee in L. specialize (AS t x). rewrite L in AS. rewrite L' in AS. simpl in AS.
|
||||
specialize (E t x). simpl in E. case: E ; first by [].
|
||||
have a:= elimT AS. specialize (a is_true_true). by rewrite a.
|
||||
Qed.
|
||||
|
||||
Lemma ltn_seq_trans T x t s : leq (class T) x t -> negb (eqop x t) -> all (leq (class T) t) s ->
|
||||
all (leq (class T) x) s && all (fun y => negb (eqop x y)) s.
|
||||
elim:s ; first by [].
|
||||
move => e s IH L N A. simpl in A. simpl. specialize (IH L N (proj2 (andP A))).
|
||||
rewrite (leq_trans (y:=t) L) ; last by rewrite -> (proj1 (andP A)). simpl.
|
||||
rewrite (leq_seq_trans L (proj2 (andP A))). simpl. rewrite (proj2 (andP IH)).
|
||||
by rewrite (ltn_trans (proj1 (andP A)) L N).
|
||||
Qed.
|
||||
|
||||
Lemma sorted_cons (T:type) (s:seq T) (x:T) : sorted (x::s) = all (leq (class T) x) s && sorted s.
|
||||
elim:s x ; first by [].
|
||||
move => t s IH x. simpl @all.
|
||||
apply trans_eq with (y:=leq (class T) x t && sorted (t::s)) ; first by []. rewrite (IH t). clear IH.
|
||||
case_eq (leq (class T) x t) ; last by []. move => xt. simpl.
|
||||
case_eq (all (leq (class T) t) s) ; last by simpl ; rewrite andbF.
|
||||
move => ts. simpl. by rewrite (leq_seq_trans xt ts).
|
||||
Qed.
|
||||
|
||||
Definition unpack K (k : forall T (c : class_of T), K T c) cT :=
|
||||
let: Pack T c _ := cT return K _ (class cT) in k _ c.
|
||||
Definition repack cT : _ -> Type -> type := let k T c p := p c in unpack k cT.
|
||||
Definition pack T c := @Pack T c T.
|
||||
|
||||
Definition eqType cT := Equality.Pack (Equality.Mixin (@eqP cT)) cT.
|
||||
|
||||
Lemma comp_ne cT x y b : comp (class cT) x y = inr unit b -> negb (@eqop cT x y).
|
||||
unfold eqop. move => e. by rewrite e.
|
||||
Qed.
|
||||
|
||||
End ClassDef.
|
||||
Module Import Exports.
|
||||
Coercion eqType : type >-> Equality.type.
|
||||
Coercion sort : type >-> Sortclass.
|
||||
Notation compType := type.
|
||||
Notation CompType := pack.
|
||||
Notation CompMixin := Mixin.
|
||||
|
||||
Canonical Structure eqType.
|
||||
Notation "[ 'compType' 'of' T 'for' cT ]" :=
|
||||
(@repack cT (@Pack T) T)
|
||||
(at level 0, format "[ 'compType' 'of' T 'for' cT ]") : form_scope.
|
||||
Notation "[ 'compType' 'of' T ]" :=
|
||||
(repack (fun c => @Pack T c) T)
|
||||
(at level 0, format "[ 'compType' 'of' T ]") : form_scope.
|
||||
End Exports.
|
||||
End Comparison.
|
||||
Export Comparison.Exports.
|
||||
|
||||
|
||||
Definition comparison (T:compType) (x y:T) := Comparison.comp (Comparison.class T) x y.
|
||||
Definition leq (T:compType) (x y:T) := Comparison.leq (Comparison.class T) x y.
|
||||
Definition sorted (T:compType) (s:seq T) := Comparison.sorted s.
|
||||
|
||||
Lemma leq_trans (T:compType) (x y z:T) : leq x y -> leq y z -> leq x z.
|
||||
move => E. by apply (Comparison.leq_trans E).
|
||||
Qed.
|
||||
|
||||
Lemma leq_refl (T:compType) (x:T) : leq x x.
|
||||
by apply (Comparison.leq_refl).
|
||||
Qed.
|
||||
|
||||
Lemma leq_seq_trans (T:compType) (x:T) t s : leq x t -> all (leq t) s -> all (leq x) s.
|
||||
move => L A. by apply (Comparison.leq_seq_trans L A).
|
||||
Qed.
|
||||
|
||||
Lemma ltn_trans (T:compType) (x y t:T) : leq x y -> leq t x -> t != x -> t != y.
|
||||
move => L L' N. by apply (Comparison.ltn_trans L L' N).
|
||||
Qed.
|
||||
|
||||
Lemma ltn_seq_trans (T:compType) (x t:T) s : leq x t -> x != t -> all (leq t) s ->
|
||||
all (leq x) s && all (fun y => x != y) s.
|
||||
move => L N A. by apply (Comparison.ltn_seq_trans L N A).
|
||||
Qed.
|
||||
|
||||
Lemma sorted_cons (T:compType) (s:seq T) (x:T) : sorted (x::s) = all (leq x) s && sorted s.
|
||||
by apply Comparison.sorted_cons.
|
||||
Qed.
|
||||
|
||||
Lemma comp_eq (T:compType) (x y:T) : (comparison x y == inl bool tt) = (x == y).
|
||||
case_eq (comparison x y == inl bool tt) => E.
|
||||
apply sym_eq. by apply E.
|
||||
case_eq (x == y) => E' ; last by [].
|
||||
case: T x y E E'. move => T. case => l c A0 A1 A2 T'. simpl.
|
||||
move => x y. have a:= (A0 x y). have A:=proj1 (A2 x y). move => F. simpl in A. rewrite F in A.
|
||||
move => e. rewrite (eqP e) in A. by rewrite (introT (A0 y y) (refl_equal _)) in A.
|
||||
Qed.
|
||||
|
||||
Lemma comp_leq (T:compType) (x y:T) : (comparison x y == inr unit true) = (leq x y && (x != y)).
|
||||
case_eq (x == y) => E. have e:=comp_eq x y. rewrite E in e.
|
||||
rewrite (eqP e). simpl. by rewrite andbF.
|
||||
|
||||
simpl. rewrite andbT. have e:(comparison x y != inl bool tt) by rewrite comp_eq ; rewrite E. clear E.
|
||||
move: T x y e.
|
||||
case => T. case => l c A0 A1 A2 T'. unfold comparison. unfold leq. simpl.
|
||||
move => x y. have a:= (proj1 (proj2 (A2 x y))). simpl in a. rewrite a. clear a.
|
||||
case_eq (l x y) => E ; last by []. simpl. case_eq (l y x) => L ; last by [].
|
||||
specialize (A2 x y). rewrite E in A2. rewrite L in A2. simpl in A2. by rewrite (proj1 A2).
|
||||
Qed.
|
||||
|
||||
Lemma comp_geq (T:compType) (x y:T) : (comparison x y == inr unit false) = (leq y x && (x != y)).
|
||||
case_eq (x == y) => E. have e:=comp_eq x y. rewrite E in e.
|
||||
rewrite (eqP e). simpl. by rewrite andbF.
|
||||
|
||||
simpl. rewrite andbT. have e:(comparison x y != inl bool tt) by rewrite comp_eq ; rewrite E. clear E.
|
||||
move: T x y e.
|
||||
case => T. case => l c A0 A1 A2 T'. unfold comparison. unfold leq. simpl.
|
||||
move => x y. have a:= (proj2 (proj2 (A2 x y))). simpl in a. rewrite a. clear a.
|
||||
case_eq (l y x) => E ; last by []. simpl. case_eq (l x y) => L ; last by [].
|
||||
specialize (A2 x y). rewrite E in A2. rewrite L in A2. simpl in A2. by rewrite (proj1 A2).
|
||||
Qed.
|
||||
|
||||
Lemma comp_neq (T:compType) (x y:T) b : comparison x y = inr _ b -> x != y.
|
||||
move => E. by apply (Comparison.comp_ne E).
|
||||
Qed.
|
||||
|
||||
Lemma comp_leqT (T:compType) (x y:T) : comparison x y = inr unit true -> leq x y.
|
||||
case:T x y => T. case => l c AS Tr CL X. simpl. move => x y. have A:= (CL x y). unfold comparison. simpl. unfold leq. simpl.
|
||||
move => E. rewrite E in A. simpl in A. have a:=proj1 (proj2 A). by case: (l x y) a.
|
||||
Qed.
|
||||
|
||||
Lemma comp_leqF (T:compType) (x y:T) : comparison x y = inr unit false -> leq y x.
|
||||
case:T x y => T. case => l c AS Tr CL X. simpl. move => x y. have A:= (CL x y). unfold comparison. simpl. unfold leq. simpl.
|
||||
move => E. rewrite E in A. simpl in A. have a:=proj2 (proj2 A). by case: (l y x) a.
|
||||
Qed.
|
||||
|
||||
Lemma leq_anti (T:compType) (x y : T) : leq x y -> leq y x -> x = y.
|
||||
unfold leq. case: T x y. move => T. case. move => le comp A B C T'. simpl.
|
||||
move => x y l l'. specialize (A x y). rewrite l in A. rewrite l' in A. simpl in A.
|
||||
by inversion A.
|
||||
Qed.
|
||||
|
||||
Fixpoint compare_nat (m n : nat) : unit + bool :=
|
||||
match m,n with
|
||||
| O,O => inl _ tt
|
||||
| S m, O => inr _ false
|
||||
| O, S n => inr _ true
|
||||
| S m, S n => compare_nat m n
|
||||
end.
|
||||
|
||||
Lemma comp_natAS : Comparison.axiomAS ssrnat.leq.
|
||||
move => x y. apply: (iffP idP) ; last by move => X ; rewrite X ; rewrite leqnn.
|
||||
move => L. by apply (anti_leq L).
|
||||
Qed.
|
||||
|
||||
Lemma nat_leqT : Comparison.axiomT ssrnat.leq.
|
||||
move => x y z. case_eq (x <= y <= z) ; last by rewrite implyFb.
|
||||
move => E. rewrite implyTb. by apply (ssrnat.leq_trans (proj1 (andP E)) (proj2 (andP E))).
|
||||
Qed.
|
||||
|
||||
Lemma comp_natCL : Comparison.axiomCL ssrnat.leq compare_nat.
|
||||
move => x y. simpl. elim: x y ; first by case.
|
||||
move => x IH. case ; first by []. move => y. simpl. specialize (IH y).
|
||||
rewrite (proj1 IH). rewrite (proj1 (proj2 IH)). by rewrite (proj2 (proj2 IH)).
|
||||
Qed.
|
||||
|
||||
Canonical Structure nat_compMixin := CompMixin (comp:=compare_nat) comp_natAS nat_leqT comp_natCL.
|
||||
Canonical Structure nat_compType := Eval hnf in CompType nat_compMixin.
|
||||
|
||||
Lemma map_map T T' T'' (f:T -> T') (g:T' -> T'') l : map g (map f l) = map (fun x => g (f x)) l.
|
||||
elim:l ; first by [].
|
||||
move => e l IH. simpl. by rewrite IH.
|
||||
Qed.
|
||||
|
||||
Section FinDom.
|
||||
Variable T:compType.
|
||||
|
||||
Section Def.
|
||||
Variable T':Type.
|
||||
|
||||
(*=FinDom *)
|
||||
Record FinDom : Type := mk_findom
|
||||
{ findom_t : seq (T * T');
|
||||
findom_P : sorted (map (@fst T T') findom_t) &&
|
||||
uniq (map (@fst T T') findom_t) }.
|
||||
(*=End *)
|
||||
Fixpoint findom_fun (f:seq (T * T')) (x:T) : option T' :=
|
||||
match f with
|
||||
| nil => None
|
||||
| (x0,y) :: r => if x == x0 then Some y else findom_fun r x
|
||||
end.
|
||||
|
||||
Definition findom_f f := findom_fun (findom_t f).
|
||||
|
||||
Definition dom f := map (@fst _ _) (findom_t f).
|
||||
Definition codom f := map (@snd _ _) (findom_t f).
|
||||
|
||||
End Def.
|
||||
|
||||
Coercion findom_f : FinDom >-> Funclass.
|
||||
|
||||
Variable T' : Type.
|
||||
|
||||
Fixpoint updpair (x:T * T') (l:seq (T * T')) : seq (T * T') :=
|
||||
match l with
|
||||
| nil => [:: x]
|
||||
| y::yr => match comparison (fst x) (fst y) with inl _ => x::yr | inr true => x::y::yr | inr false => y::updpair x yr end
|
||||
end.
|
||||
|
||||
Lemma all_leq_upd x (s:seq (T * T')) t : all (leq t) (map (@fst _ _) s) -> leq t x.1 ->
|
||||
all (leq t) (map (@fst _ _) (updpair x s)).
|
||||
elim:s t ; first simpl. move => t. by rewrite andbT.
|
||||
case => t0 t0' s IH t. simpl. move => A L.
|
||||
case: (comparison x.1 t0) (comp_eq x.1 t0) (comp_leq x.1 t0) (comp_geq x.1 t0) ; case ; rewrite eq_refl.
|
||||
- move => B _ _. rewrite (eqP (sym_eq B)) in L. simpl. rewrite (proj2 (andP A)).
|
||||
rewrite <- (eqP (sym_eq B)) in L. by rewrite L.
|
||||
- move => _ B _. simpl. rewrite L. simpl. by apply A.
|
||||
- move => _ _ B. simpl. rewrite (proj1 (andP A)). simpl. apply IH ; first by apply (proj2 (andP A)).
|
||||
by apply L.
|
||||
Qed.
|
||||
|
||||
Lemma notin_updpair x s t : t != x.1 -> t \notin map (@fst _ _) s -> t \notin map (@fst _ _) (updpair x s).
|
||||
elim: s t. simpl. move => t e _. rewrite in_cons. by case: (t == x.1) e.
|
||||
case => t t' s IH t0 e. simpl. rewrite in_cons.
|
||||
case: (comparison x.1 t) (comp_eq x.1 t) (comp_leq x.1 t) (comp_geq x.1 t) ; case ; rewrite eq_refl.
|
||||
- move => E. rewrite (eqP (sym_eq E)) in e. simpl. rewrite in_cons.
|
||||
by rewrite (eqP (sym_eq E)).
|
||||
- move => _ A _. simpl. do 2 rewrite in_cons.
|
||||
case: (t0 == t) ; first by []. simpl. by case: (t0 == x.1) e.
|
||||
- move=> _ _ A. simpl. rewrite in_cons. case: (t0 == t) ; first by [].
|
||||
simpl. move => X. by apply IH.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma findom_fun_notin T0 t s : t \notin map (@fst _ T0) s -> findom_fun s t = None.
|
||||
elim: s ; first by [].
|
||||
case => e2 e3 s IH. simpl. rewrite in_cons. rewrite negb_or. move => P.
|
||||
specialize (IH (proj2 (andP P))). rewrite IH. rewrite <- if_neg. by rewrite (proj1 (andP P)).
|
||||
Qed.
|
||||
|
||||
Lemma findom_upd (a:T') l s : findom_fun (updpair (l, a) s) l = Some a.
|
||||
elim: s. simpl. by rewrite eq_refl.
|
||||
- case => e0 e1 s IH. simpl. case_eq (comparison l e0) ; case => E.
|
||||
+ simpl. by rewrite eq_refl.
|
||||
+ simpl. by rewrite eq_refl.
|
||||
+ simpl. have b:=comp_geq l e0. rewrite E in b. rewrite eq_refl in b. have c:=sym_eq b.
|
||||
have N:=proj2 (andP c). rewrite <- if_neg. rewrite N. by apply IH.
|
||||
Qed.
|
||||
|
||||
Lemma findom_upd2 (a:T') x l s : x != l -> findom_fun (updpair (l, a) s) x = findom_fun s x.
|
||||
move => P. elim: s. simpl. rewrite <- if_neg. by rewrite P.
|
||||
case => e0 e1 s IH. simpl. case_eq (comparison l e0) ; case => E.
|
||||
- have e:=comp_eq l e0. rewrite E in e. rewrite eq_refl in e. simpl. rewrite <- (eqP (sym_eq e)).
|
||||
rewrite <- if_neg. rewrite P. rewrite <- if_neg. by rewrite P.
|
||||
- simpl. rewrite <- if_neg. by rewrite P.
|
||||
- simpl. case_eq (x == e0) ; first by []. move => E'. by apply IH.
|
||||
Qed.
|
||||
|
||||
Lemma all_diff_notin (t:T) s : all (fun y : T => t != y) s = (t \notin s).
|
||||
elim: s ; first by [].
|
||||
move => t0 s IH. simpl. rewrite in_cons. rewrite negb_or. case_eq (t != t0) ; last by []. move => e. by apply IH.
|
||||
Qed.
|
||||
|
||||
Lemma updpairP (x:T * T') f : sorted (T:=T) (map (@fst _ T') (updpair x (findom_t f))) &&
|
||||
uniq (map (@fst _ _) (updpair x (findom_t f))).
|
||||
case: f. elim ; first by [].
|
||||
case => t t' s IH. simpl @map. rewrite sorted_cons.
|
||||
case: (comparison x.1 t) (comp_eq x.1 t) (comp_leq x.1 t) (comp_geq x.1 t) ; case ; rewrite eq_refl.
|
||||
- move => A. simpl @map. rewrite sorted_cons. simpl. by rewrite (eqP (sym_eq A)).
|
||||
- move => _ A _. simpl @map. do 2 rewrite sorted_cons. simpl. move => B.
|
||||
rewrite (proj2 (andP B)). rewrite (proj1 (andP B)). rewrite (proj1 (andP (sym_eq A))). simpl.
|
||||
have C:=(ltn_seq_trans (proj1 (andP (sym_eq A))) (proj2 (andP (sym_eq A))) (proj1 (andP (proj1 (andP B))))).
|
||||
rewrite (proj1 (andP C)). rewrite -> all_diff_notin in C. rewrite in_cons.
|
||||
case: (x.1 == t) (proj2 (andP (sym_eq A))) ; first by []. simpl.
|
||||
by rewrite (proj2 (andP C)).
|
||||
- move => _ _ A. simpl @map. rewrite sorted_cons. simpl. move => B.
|
||||
have Y:sorted (map (@fst _ _) s) && uniq (map (@fst _ _) s).
|
||||
rewrite (proj2 (andP (proj1 (andP B)))). by rewrite (proj2 (andP (proj2 (andP B)))).
|
||||
specialize (IH Y). simpl in IH. rewrite (proj1 (andP IH)). rewrite (proj2 (andP IH)).
|
||||
do 2 rewrite andbT. rewrite (all_leq_upd (proj1 (andP (proj1 (andP B)))) (proj1 (andP (sym_eq A)))).
|
||||
simpl. apply notin_updpair ; last by apply (proj1 (andP (proj2 (andP B)))).
|
||||
rewrite eq_sym. by apply (proj2 (andP (sym_eq A))).
|
||||
Qed.
|
||||
|
||||
Definition updMap (t:T) (t':T') (f:FinDom T') : FinDom T' := mk_findom (@updpairP (t,t') f).
|
||||
|
||||
Lemma updMap_simpl t t' f : updMap t t' f t = Some t'.
|
||||
case: f => s f. unfold findom_f. simpl. clear f. elim:s ; first by simpl ; rewrite eq_refl.
|
||||
move => e s IH. simpl. case_eq (comparison t e.1) ; case.
|
||||
+ move => eq. rewrite eq. simpl. by rewrite -> eq_refl.
|
||||
+ move => eq ; rewrite eq. simpl. by rewrite eq_refl.
|
||||
+ move => eq ; rewrite eq. simpl. case: e eq. move => e0 e1. rewrite IH. simpl. move => X. have a:=comp_geq t e0. rewrite X in a.
|
||||
rewrite eq_refl in a. have b:= sym_eq a. have f:=(proj2 (andP b)). case_eq (t == e0) => E ; first by rewrite E in f.
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Lemma updMap_simpl2 t t0 t' f : t0 != t -> updMap t t' f t0 = f t0.
|
||||
case:f => s f. unfold findom_f. simpl. clear f. move => ne. elim: s.
|
||||
- simpl. case_eq (t0 == t) =>e ; last by []. by rewrite e in ne.
|
||||
- move => e s IH. simpl. case_eq (comparison t e.1) ; case.
|
||||
+ move => E. rewrite E. simpl. rewrite <- if_neg. rewrite ne. have te:=comp_eq t e.1. rewrite E in te. rewrite eq_refl in te.
|
||||
have te':=sym_eq te. rewrite (eqP te') in ne. clear E te te'. case: e ne. simpl. move => e0 e1 e. rewrite <- if_neg.
|
||||
by rewrite e.
|
||||
+ move => eq ; rewrite eq. simpl. rewrite <- if_neg. by rewrite ne.
|
||||
+ move => eq ; rewrite eq. simpl. by rewrite IH.
|
||||
Qed.
|
||||
|
||||
Lemma NoneP : sorted (T:=T) (map (@fst T T') [::]) && uniq (map (@fst T T') [::]).
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition fdEmpty : FinDom T' := mk_findom NoneP.
|
||||
|
||||
Lemma findom_ind (f:FinDom T') (P:FinDom T' -> Prop) : P fdEmpty ->
|
||||
(forall f n x, all (leq n) (dom f) -> n \notin dom f -> P f -> P (updMap n x f)) -> P f.
|
||||
move => E C. case: f. elim.
|
||||
- simpl. move => X. unfold fdEmpty in E. rewrite -> (eq_irrelevance NoneP X) in E. by apply E.
|
||||
- case => a b s. move => IH X.
|
||||
have X':=X. simpl @map in X'. rewrite sorted_cons in X'. simpl in X'.
|
||||
have Y:sorted (map (@fst _ _) s) && uniq (map (@fst _ _) s).
|
||||
rewrite -> (proj2 (andP (proj2 (andP X')))). by rewrite -> (proj2 (andP (proj1 (andP X')))).
|
||||
specialize (IH Y). specialize (C (mk_findom Y) a b).
|
||||
have d:a \notin dom (mk_findom Y). unfold dom. simpl. simpl in X. by apply (proj1 (andP (proj2 (andP X)))).
|
||||
specialize (C (proj1 (andP (proj1 (andP X')))) d IH).
|
||||
have e:(updMap a b (mk_findom Y)) = (mk_findom X). simpl @map in X.
|
||||
unfold updMap. have A:= (updpairP (a, b) (mk_findom Y)).
|
||||
rewrite (eq_irrelevance (updpairP (a, b) (mk_findom Y)) A).
|
||||
simpl in A.
|
||||
have ee:(updpair (a, b) s) = (a,b)::s. clear Y C IH X d. case: s X' A ; first by [].
|
||||
case => a0 b0 s. simpl. move => A B. have L:= comp_leq a a0.
|
||||
rewrite (proj1 (andP (proj1 (andP (proj1 (andP A)))))) in L. simpl in L.
|
||||
have ee:=(proj1 (andP (proj2 (andP A)))). rewrite in_cons in ee.
|
||||
have x:(a != a0). by case: (a == a0) ee. rewrite x in L. clear ee x. by rewrite (eqP L).
|
||||
move: A. rewrite ee. simpl @map. move => A. by rewrite (eq_irrelevance X A).
|
||||
rewrite <- e. by apply C.
|
||||
Qed.
|
||||
|
||||
Lemma dom_fdEmpty : dom fdEmpty = [::].
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Lemma dom_empty_eq f : dom f = [::] -> f = fdEmpty.
|
||||
unfold dom. case: f. simpl. case ; last by [].
|
||||
simpl. move => X _. by rewrite (eq_irrelevance X NoneP).
|
||||
Qed.
|
||||
|
||||
Fixpoint rempair (x:T) (l:seq (T * T')) : seq (T * T') :=
|
||||
match l with
|
||||
| nil => [::]
|
||||
| y::yr => match comparison x (fst y) with inl _ => yr | inr true => y::yr | inr false => y::rempair x yr end
|
||||
end.
|
||||
|
||||
Lemma remMapP t f : sorted (T:=T) (map (@fst _ _) (rempair t (findom_t f))) &&
|
||||
uniq (map (@fst _ _) (rempair t (findom_t f))).
|
||||
case: f => s P. simpl. elim: s P ; first by [].
|
||||
case. move => t0 t0' s IH P. simpl.
|
||||
case_eq (comparison t t0) ; case.
|
||||
- move => C. simpl @map in P. rewrite sorted_cons in P. rewrite (proj2 (andP (proj1 (andP P)))). simpl.
|
||||
simpl in P. by rewrite (proj2 (andP (proj2 (andP P)))).
|
||||
- by [].
|
||||
- move => C. simpl @map. simpl @map in P. rewrite sorted_cons in P. simpl in P.
|
||||
rewrite (proj2 (andP (proj1 (andP P)))) in IH. rewrite (proj2 (andP (proj2 (andP P)))) in IH.
|
||||
specialize (IH is_true_true). rewrite sorted_cons. rewrite (proj1 (andP IH)). simpl. rewrite (proj2 (andP IH)).
|
||||
do 2 rewrite andbT. clear IH.
|
||||
have a:=(proj1 (andP (proj1 (andP P)))). have b:=(proj1 (andP (proj2 (andP P)))). clear C P.
|
||||
elim: s a b ; first by []. case => e0 e1 s IH P Q. simpl. simpl in P. simpl in Q.
|
||||
case_eq (comparison t e0) ; case.
|
||||
+ rewrite in_cons in Q. rewrite negb_or in Q. rewrite (proj2 (andP Q)). by rewrite (proj2 (andP P)).
|
||||
+ simpl. rewrite Q. by rewrite P.
|
||||
+ simpl. specialize (IH (proj2 (andP P))). rewrite in_cons in Q. rewrite negb_or in Q. specialize (IH (proj2 (andP Q))).
|
||||
rewrite in_cons. rewrite negb_or. rewrite (proj1 (andP IH)). rewrite (proj2 (andP IH)).
|
||||
rewrite (proj1 (andP Q)). by rewrite (proj1 (andP P)).
|
||||
Qed.
|
||||
|
||||
Definition remMap (t:T) (f:FinDom T') : FinDom T' := mk_findom (@remMapP t f).
|
||||
|
||||
Lemma updpair_least n x s : all (leq n) (map (@fst _ _) s) -> n \notin (map (@fst _ _) s) ->
|
||||
(updpair (n, x) s) = (n,x)::s.
|
||||
case:s ; first by [].
|
||||
case => t t' s X Y. simpl. simpl in X. simpl in Y. rewrite in_cons in Y.
|
||||
case: (comparison n t) (comp_eq n t) (comp_leq n t) (comp_geq n t) ; case ; rewrite eq_refl.
|
||||
- move => A _ _. by rewrite (sym_eq A) in Y.
|
||||
- by [].
|
||||
- move => _ _ A. have e:=leq_anti (proj1 (andP X)) (proj1 (andP (sym_eq A))). rewrite e in A.
|
||||
rewrite eq_refl in A. by rewrite andbF in A.
|
||||
Qed.
|
||||
|
||||
Lemma seq_ext_nil (g:seq T) : [::] =i g -> g = [::].
|
||||
case: g ; first by [].
|
||||
move => s r X. specialize (X s). rewrite in_nil in X. rewrite in_cons in X. by rewrite eq_refl in X.
|
||||
Qed.
|
||||
|
||||
Lemma findom_ext (f g:FinDom T') : dom f =i dom g -> (forall x, x \in dom g -> f x = g x) -> g = f.
|
||||
case: f g. unfold findom_f. unfold dom. simpl. move => f Pf. case. simpl. move => g Pg D C.
|
||||
have e:f = g. elim: f g Pf Pg C D. simpl. move => g _ e X a. move: (seq_ext_nil a). clear. by case: g.
|
||||
case => c e f IH. case. simpl. move => _ _ _ F. specialize (F c). rewrite in_nil in F.
|
||||
by rewrite in_cons in F ; simpl in F ; rewrite eq_refl in F.
|
||||
case => c' e' g. simpl @map. do 2 rewrite sorted_cons.
|
||||
move => A B C X.
|
||||
have Y:= X c. have XX:=X. specialize (X c'). do 2 rewrite in_cons in X. do 2 rewrite in_cons in Y.
|
||||
rewrite (eq_sym c' c) in X. rewrite eq_refl in X. rewrite eq_refl in Y. simpl in X. simpl in Y.
|
||||
have F:c != c' -> False. simpl in A,B.
|
||||
have Z:c != c' -> c = c'.
|
||||
case: (c == c') X Y ; first by []. simpl. move => X Y _.
|
||||
have l:= allP (proj1 (andP (proj1 (andP A)))) _ X.
|
||||
have l':= allP (proj1 (andP (proj1 (andP B)))) _ (sym_eq Y).
|
||||
by apply (leq_anti l l'). move => F. specialize (Z F). rewrite Z in F. by rewrite eq_refl in F.
|
||||
have e0:(c == c') ; case: (c == c') F ; first by []. move => Z. by case (Z is_true_true).
|
||||
move => _. rewrite <- (eqP e0).
|
||||
have C':=C c. rewrite in_cons in C'. rewrite e0 in C'. specialize (C' is_true_true).
|
||||
simpl in C'. rewrite eq_refl in C'. rewrite e0 in C'. case: C' => C'. rewrite <- C'. f_equal.
|
||||
apply IH. simpl in A. rewrite (proj2 (andP (proj2 (andP A)))). by rewrite (proj2 (andP (proj1 (andP A)))).
|
||||
simpl in B. rewrite (proj2 (andP (proj2 (andP B)))). by rewrite (proj2 (andP (proj1 (andP B)))).
|
||||
move => x ix.
|
||||
specialize (C x). rewrite in_cons in C. rewrite ix in C. rewrite orbT in C. specialize (C is_true_true).
|
||||
simpl in C. simpl in B. have F:x == c -> False. move => F. rewrite (eqP F) in ix. rewrite (eqP e0) in ix.
|
||||
rewrite ix in B. by rewrite andbF in B.
|
||||
rewrite <- (eqP e0) in C. case: (x == c) F C ; last by [].
|
||||
move => F. by case: (F is_true_true).
|
||||
move => a. specialize (XX a). rewrite <- (eqP e0) in XX.
|
||||
do 2 rewrite in_cons in XX. rewrite <- (eqP e0) in B. simpl in A. simpl in B.
|
||||
case_eq (a == c) => aa. rewrite (eqP aa). move: (proj1 (andP (proj2 (andP A)))).
|
||||
move: (proj1 (andP (proj2 (andP B)))). case: (c \in map (@fst _ _) f) ; first by [].
|
||||
case: (c \in map (@fst _ _) g) ; by [].
|
||||
rewrite aa in XX. simpl in XX. by apply XX.
|
||||
move => F. case: (F is_true_true).
|
||||
|
||||
move: C D Pg Pf. rewrite e. clear f e. move => C D Pg Pf. by rewrite (eq_irrelevance Pg Pf).
|
||||
Qed.
|
||||
|
||||
Lemma in_filter (Te:eqType) P l (x:Te) : x \in filter P l = (x \in l) && P x.
|
||||
elim:l x ; first by [].
|
||||
move => t s IH x. simpl. rewrite in_cons. have e:x == t -> P t = P x. move => E. by rewrite (eqP E).
|
||||
case: (P t) e.
|
||||
- rewrite in_cons. case: (x == t) ; last move => F. simpl. move => X. by apply (X is_true_true).
|
||||
simpl. by apply IH.
|
||||
- case: (x == t) ; simpl. move => X. specialize (X is_true_true). specialize (IH x). rewrite <- X in IH.
|
||||
rewrite andbF in IH. rewrite IH. by rewrite <- X.
|
||||
move => A. by apply IH.
|
||||
Qed.
|
||||
|
||||
Lemma indomEmpty i : (i \in dom fdEmpty) = false.
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Lemma indomUpdMap i f x y : i \in dom (updMap x y f) = (i == x) || (i \in dom f).
|
||||
move: x y. apply (findom_ind f).
|
||||
- move => x y. rewrite indomEmpty. rewrite orbF. unfold updMap. unfold fdEmpty. unfold dom. simpl. rewrite in_cons.
|
||||
rewrite in_nil. by rewrite orbF.
|
||||
- clear f. move => f n x A I IH x' y'. unfold dom. simpl. rewrite (updpair_least _ A).
|
||||
simpl. rewrite in_cons. case: (comparison x' n) (comp_eq x' n) (comp_leq x' n) (comp_geq x' n) ; case ; rewrite eq_refl.
|
||||
+ move => B _ _. simpl. rewrite in_cons. rewrite (eqP (sym_eq B)). clear x' B y'.
|
||||
by case: (i == n).
|
||||
+ move => _ B _. simpl. rewrite in_cons. by rewrite in_cons.
|
||||
+ move => _ _ B. simpl. rewrite in_cons. specialize (IH x' y'). unfold dom in IH. simpl in IH. rewrite IH.
|
||||
case: (i == n) ; simpl. by rewrite orbT.
|
||||
+ by [].
|
||||
by apply I.
|
||||
Qed.
|
||||
|
||||
Lemma updMapCom f x y x' y' : x != x' -> updMap x y (updMap x' y' f) = updMap x' y' (updMap x y f).
|
||||
move => e. apply findom_ext.
|
||||
- move => a. do 4 rewrite indomUpdMap.
|
||||
case: (a == x') ; simpl ; first by rewrite orbT. by [].
|
||||
- move => n I. do 2 rewrite indomUpdMap in I.
|
||||
case_eq (n == x') => F.
|
||||
+ rewrite (eqP F). rewrite updMap_simpl. rewrite eq_sym in e.
|
||||
rewrite (updMap_simpl2 y _ e). by rewrite updMap_simpl.
|
||||
+ have ee:n != x' by case: (n == x') F. clear F.
|
||||
rewrite (updMap_simpl2 y' _ ee). case_eq (x == n).
|
||||
* move => e0. rewrite (eqP e0). by do 2 rewrite updMap_simpl.
|
||||
* move => e0. have e1:x != n by case: (x == n) e0.
|
||||
clear e0. rewrite eq_sym in e1. do 2 rewrite (updMap_simpl2 y _ e1).
|
||||
by rewrite (updMap_simpl2 y' _ ee).
|
||||
Qed.
|
||||
|
||||
Lemma indomRemMap a f t : (a \in dom (remMap t f)) = (a != t) && (a \in dom f).
|
||||
move: t ; apply (findom_ind f) ; clear f.
|
||||
- move => t. rewrite indomEmpty. by rewrite andbF.
|
||||
- move => f n x A I IH t.
|
||||
rewrite indomUpdMap. unfold dom. simpl. rewrite (updpair_least _ A) ; last by apply I.
|
||||
simpl. case: (comparison t n) (comp_eq t n) (comp_leq t n) (comp_geq t n) ; case ; rewrite eq_refl.
|
||||
+ move => B _ _. rewrite (eqP (sym_eq B)). clear t B.
|
||||
case_eq (a == n) ; last by []. simpl. move => e. rewrite <- (eqP e) in I. unfold dom in I.
|
||||
by case: (a \in map (@fst _ _) (findom_t f)) I.
|
||||
+ move => _ B _. simpl. rewrite in_cons.
|
||||
case_eq (a == n) => e. simpl. rewrite eq_sym. rewrite (eqP e). by rewrite (proj2 (andP (sym_eq B))).
|
||||
simpl. case_eq (a == t) ; last by [].
|
||||
move => e'. rewrite (eqP e'). rewrite -> (eqP e') in IH, e. clear e' a.
|
||||
rewrite e in B. simpl in B. rewrite andbT in B. have aa:=ltn_seq_trans (sym_eq B) _ A.
|
||||
rewrite e in aa. simpl in aa. specialize (aa is_true_true).
|
||||
have bb:=proj2 (andP aa). rewrite all_diff_notin in bb. simpl. unfold dom in bb.
|
||||
by case: (t \in map (@fst _ _) (findom_t f)) bb.
|
||||
+ move => _ _ B. simpl. rewrite in_cons. rewrite IH. case_eq (a == n) ; simpl ; last by [].
|
||||
move => e. rewrite (eqP e). rewrite eq_sym. by rewrite (proj2 (andP (sym_eq B))).
|
||||
Qed.
|
||||
|
||||
Lemma remMap_simpl t (f:FinDom T') : remMap t f t = None.
|
||||
case: f => s P. unfold findom_f. simpl.
|
||||
elim: s P ; first by [].
|
||||
case => e0 e1 s IH. simpl @map. rewrite sorted_cons. move => P. simpl.
|
||||
case_eq (comparison t e0) ;case.
|
||||
- have Q:=proj1 (andP (proj2 (andP P))). clear IH P. move => C. have a:=comp_eq t e0. rewrite C in a. rewrite eq_refl in a.
|
||||
have b:=sym_eq a. rewrite <- (eqP b) in Q. clear a b C. by rewrite (findom_fun_notin Q).
|
||||
- simpl. move => C. have a:=comp_leq t e0. rewrite C in a. rewrite eq_refl in a. have b:=sym_eq a.
|
||||
rewrite <- if_neg. rewrite (proj2 (andP b)).
|
||||
have c:=ltn_seq_trans (proj1 (andP b)) (proj2 (andP b)) (proj1 (andP (proj1 (andP P)))).
|
||||
rewrite ((all_diff_notin _ _)) in c. by rewrite (findom_fun_notin (proj2 (andP c))).
|
||||
- move => C. simpl in P. rewrite (proj2 (andP (proj1 (andP P)))) in IH. rewrite (proj2 (andP (proj2 (andP P)))) in IH.
|
||||
simpl. rewrite IH ; last by []. rewrite <- if_neg. have a:=comp_geq t e0. rewrite C in a. rewrite eq_refl in a.
|
||||
have b:=sym_eq a. by rewrite (proj2 (andP b)).
|
||||
Qed.
|
||||
|
||||
Lemma remMap_simpl2 t t' (f:FinDom T') : t' != t -> remMap t f t' = f t'.
|
||||
case:f => s P. unfold findom_f. simpl. move => E.
|
||||
elim: s P ; first by [].
|
||||
case => e0 e1 s IH. simpl @map. rewrite sorted_cons. move => P. simpl.
|
||||
case_eq (comparison t e0) ;case.
|
||||
- have Q:=proj1 (andP (proj2 (andP P))). clear IH P. move => C. have a:=comp_eq t e0. rewrite C in a. rewrite eq_refl in a.
|
||||
have b:=sym_eq a. rewrite <- (eqP b) in Q. rewrite <- (eqP b). rewrite <- if_neg. by rewrite E.
|
||||
- by [].
|
||||
- move => C. simpl in P. simpl. rewrite (proj2 (andP (proj1 (andP P)))) in IH. rewrite (proj2 (andP (proj2 (andP P)))) in IH.
|
||||
simpl. by rewrite IH.
|
||||
Qed.
|
||||
|
||||
Lemma remUpdMap f n t x : t != n -> remMap t (updMap n x f) = updMap n x (remMap t f).
|
||||
move => e. apply findom_ext.
|
||||
- move => a. rewrite indomUpdMap. do 2 rewrite indomRemMap. rewrite indomUpdMap.
|
||||
case_eq (a == n) => ee ; last by []. simpl. rewrite <- (eqP ee) in e. rewrite eq_sym. by rewrite e.
|
||||
- move => a X. rewrite indomRemMap in X. rewrite indomUpdMap in X.
|
||||
case_eq (a == n) => ee. rewrite (eqP ee). rewrite updMap_simpl. rewrite (eqP ee) in X.
|
||||
rewrite (remMap_simpl2 _ (proj1 (andP X))). by rewrite updMap_simpl.
|
||||
have en:a != n by case: (a == n) ee. clear ee.
|
||||
rewrite (updMap_simpl2 _ _ en). do 2 rewrite (remMap_simpl2 _ (proj1 (andP X))).
|
||||
by rewrite (updMap_simpl2 _ _ en).
|
||||
Qed.
|
||||
|
||||
Section Eq.
|
||||
Variable Teq : eqType.
|
||||
Definition findom_eq (f f':FinDom Teq) : bool := findom_t f == findom_t f'.
|
||||
|
||||
Lemma findom_eqP : Equality.axiom findom_eq.
|
||||
move => f f'. apply: (iffP idP) ; last by move => e ; rewrite e ; apply eq_refl.
|
||||
unfold findom_eq. case: f. move => l s. case: f'. simpl. move => l' s'.
|
||||
move => X. move: s. rewrite (eqP X). clear. move => s. by rewrite (eq_irrelevance s s').
|
||||
Qed.
|
||||
|
||||
Definition findom_leq (f f':FinDom Teq) : bool := all (fun x => f x == f' x) (dom f).
|
||||
|
||||
Lemma findom_leq_refl (f:FinDom Teq) : findom_leq f f.
|
||||
case: f. unfold findom_leq. unfold dom. unfold findom_f. simpl. move => s P.
|
||||
by apply (introT (@allP T _ (map (@fst _ _) s))).
|
||||
Qed.
|
||||
|
||||
End Eq.
|
||||
|
||||
Lemma map_pmap X Y (f:X -> Y) Z (g:Z -> option X) s : map f (pmap g s) = pmap (fun x => option_map f (g x)) s.
|
||||
elim: s ; first by [].
|
||||
move => t s IH. simpl. case_eq (g t). move => gt e. simpl. by rewrite IH. simpl. by rewrite IH.
|
||||
Qed.
|
||||
|
||||
Lemma all_map_filter X Y (f:X -> Y) p g l : all p (map f l) -> all p (map f (filter g l)).
|
||||
elim: l ; first by []. move => t s IH P.
|
||||
simpl. simpl in P. specialize (IH (proj2 (andP P))). case_eq (g t) ; last by [].
|
||||
move => e. simpl. rewrite IH. by rewrite (proj1 (andP P)).
|
||||
Qed.
|
||||
|
||||
Lemma sorted_map_filter X (f:X -> T) g l : sorted (map f l) -> sorted (T:=T) (map f (filter g l)).
|
||||
elim: l ; first by []. move => t s IH. simpl @map. rewrite sorted_cons. move => P.
|
||||
specialize (IH (proj2 (andP P))).
|
||||
case (g t) ; last by []. simpl @map ; rewrite sorted_cons. rewrite IH. by rewrite (all_map_filter _ (proj1 (andP P))).
|
||||
Qed.
|
||||
|
||||
Lemma notin_map_filter X (Y:eqType) (f:X -> Y) p s x :x \notin map f s -> (x \notin map f (filter p s)).
|
||||
elim: s ; first by [].
|
||||
move => t s IH. simpl. rewrite in_cons. rewrite negb_or. move => P. specialize (IH (proj2 (andP P))).
|
||||
case (p t) ; last by []. simpl. rewrite in_cons ; rewrite negb_or. rewrite IH. by rewrite (proj1 (andP P)).
|
||||
Qed.
|
||||
|
||||
Lemma uniq_map_filter X (Y:eqType) (f:X -> Y) p l : uniq (map f l) -> uniq (map f (filter p l)).
|
||||
elim: l ; first by []. move => t s IH. simpl. move => P. specialize (IH (proj2 (andP P))).
|
||||
case (p t) ; last by []. simpl. rewrite IH. by rewrite (notin_map_filter _ (proj1 (andP P))).
|
||||
Qed.
|
||||
|
||||
Lemma post_compP T0 (g : T' -> option T0) (f:FinDom T') : sorted (T:=T)
|
||||
(map (@fst _ _) (pmap (fun p : T * T' => option_map (fun r => (p.1,r)) (g p.2)) (findom_t f))) &&
|
||||
uniq (map (@fst _ _) (pmap (fun p : T * T' => option_map (fun r => (p.1,r)) (g p.2)) (findom_t f))).
|
||||
case: f. simpl. move => l. rewrite map_pmap.
|
||||
have e:pmap (fun x => option_map (@fst _ _) (option_map [eta pair x.1] (g x.2))) l = (map (@fst _ T') (filter (fun p => g p.2) l)). elim: l ; first by []. case => e0 e1 s IH. simpl. case (g e1) ; simpl ; by rewrite IH.
|
||||
rewrite e. clear e. move => P. rewrite (sorted_map_filter _ (proj1 (andP P))). simpl.
|
||||
by rewrite (uniq_map_filter _ (proj2 (andP P))).
|
||||
Qed.
|
||||
|
||||
Definition post_comp T0 (g : T' -> option T0) (f:FinDom T') : FinDom T0 := mk_findom (post_compP g f).
|
||||
|
||||
Definition option_bind Y Z (g:Y -> option Z) (y:option Y) : option Z :=
|
||||
match y with | None => None | Some fx => g fx end.
|
||||
|
||||
Lemma in_pmap (A B : eqType) (f:A -> option B) x s y : Some x = f y -> y \in s -> x \in pmap f s.
|
||||
elim: s y. simpl. move => y. by do 2 rewrite in_nil.
|
||||
move => a s IH y. simpl. rewrite in_cons.
|
||||
move => E. specialize (IH _ E). case_eq (y == a) => e. rewrite (eqP e) in E. rewrite <- E. simpl. rewrite in_cons.
|
||||
by rewrite eq_refl.
|
||||
simpl. move => X. specialize (IH X). case: (f a) ; simpl ; last by [].
|
||||
move => aa. rewrite in_cons. rewrite IH. by rewrite orbT.
|
||||
Qed.
|
||||
|
||||
Lemma notin_pmap (A B : eqType) (f:A -> option B) x s : (forall y, Some x = f y -> y \notin s) -> x \notin pmap f s.
|
||||
elim: s. simpl. by [].
|
||||
move => t s IH C. simpl. case_eq (f t) => e ; simpl.
|
||||
have C':=C.
|
||||
specialize (C t). move => ee. rewrite ee in C. rewrite in_cons. case_eq (x == e) => aa.
|
||||
rewrite (eqP aa) in C. specialize (C (refl_equal _)). rewrite in_cons in C. rewrite eq_refl in C. by [].
|
||||
simpl. apply IH. move => y X. specialize (C' _ X). rewrite in_cons in C'. by case: (y == t) C'.
|
||||
apply IH. move => y E. specialize (C y E). rewrite in_cons in C. by case: (y == t) C.
|
||||
Qed.
|
||||
|
||||
Lemma post_comp_simpl T0 (g : T' -> option T0) (f:FinDom T') t : post_comp g f t = option_bind g (f t).
|
||||
apply (findom_ind f) ; clear f ; first by [].
|
||||
move => f n x A I X. unfold post_comp. unfold findom_f. simpl. rewrite (updpair_least _ A I). simpl.
|
||||
case_eq (t == n) => e. rewrite (eqP e). rewrite (eqP e) in X. clear e t. simpl. case: (g x). simpl. by rewrite eq_refl.
|
||||
simpl. rewrite findom_fun_notin ; first by []. rewrite map_pmap.
|
||||
have ll:forall s, n \notin (map (@fst _ _) s) -> n \notin pmap (fun x0 : T * T' => option_map (@fst _ _) (option_map [eta pair x0.1] (g x0.2))) s.
|
||||
elim ; first by [].
|
||||
case => t0 t1 s IH. simpl. rewrite in_cons. case (g t1). simpl. rewrite in_cons. by case: (n == t0).
|
||||
simpl. move => aa. apply IH. by case: (n == t0) aa.
|
||||
rewrite ll ; first by []. clear ll. by apply I.
|
||||
case_eq (f t) => ft. move => ee. unfold findom_f in ee. rewrite ee. simpl.
|
||||
case: (g x). simpl. rewrite e. move => _. unfold findom_f in X. rewrite ee in X. simpl in X.
|
||||
by rewrite <- X.
|
||||
simpl. unfold findom_f in X. rewrite ee in X. simpl in X. by rewrite <- X.
|
||||
unfold findom_f in ft. rewrite ft. simpl. case: (g x) ; simpl. rewrite e.
|
||||
move => _. unfold findom_f in X. rewrite ft in X. simpl in X. by apply X.
|
||||
unfold findom_f in X. rewrite ft in X. by apply X.
|
||||
Qed.
|
||||
|
||||
Lemma dom_post_comp T0 (g : T' -> option T0) (f:FinDom T') : dom (post_comp g f) =
|
||||
filter (fun x => isSome (option_bind g (f x))) (dom f).
|
||||
case: f. unfold findom_f. unfold dom. simpl. move => s Ps. have U:=proj2 (andP Ps). clear Ps.
|
||||
elim: s U ; first by [].
|
||||
case => c e f. simpl. move => IH. rewrite eq_refl. simpl. case_eq (g e).
|
||||
- move => ge E. simpl. move => X.
|
||||
rewrite IH ; last by rewrite (proj2 (andP X)). f_equal. apply: eq_in_filter => x I.
|
||||
case_eq (x == c) => E' ; last by []. by rewrite (eqP E') in I ; rewrite I in X.
|
||||
- move => E. simpl. move => X.
|
||||
rewrite IH ; last by rewrite (proj2 (andP X)). apply: eq_in_filter => x I.
|
||||
case_eq (x == c) => E' ; last by []. by rewrite (eqP E') in I ; rewrite I in X.
|
||||
Qed.
|
||||
|
||||
Lemma findom_undef (f:FinDom T') x (P:x \notin dom f) : f x = None.
|
||||
case:f x P => s P x C. unfold dom in C. simpl in C. by apply findom_fun_notin.
|
||||
Qed.
|
||||
|
||||
Lemma findom_indom (f:FinDom T') t : (t \in dom f) = isSome (f t).
|
||||
case: f. unfold dom. unfold findom_f. simpl. move => s P.
|
||||
elim: s P ; first by [].
|
||||
case => t0 t0' s IH P. simpl. rewrite in_cons. case_eq (t == t0) => E ; first by [].
|
||||
- simpl. apply IH. simpl @map in P. rewrite sorted_cons in P. simpl in P.
|
||||
rewrite (proj2 (andP (proj1 (andP P)))). by rewrite (proj2 (andP (proj2 (andP P)))).
|
||||
Qed.
|
||||
|
||||
Lemma filter_some T0 (s:seq T0) : filter some s = s.
|
||||
elim: s ; first by [].
|
||||
move => t s IH. simpl. by rewrite IH.
|
||||
Qed.
|
||||
|
||||
Lemma dom_post_compS T0 (g : T' -> T0) (f:FinDom T') : dom (post_comp (fun x => Some (g x)) f) = dom f.
|
||||
rewrite dom_post_comp. apply trans_eq with (y:=filter [eta some] (dom f)) ; last by rewrite filter_some.
|
||||
apply: eq_in_filter => x I. rewrite findom_indom in I. by case: (f x) I.
|
||||
Qed.
|
||||
|
||||
End FinDom.
|
||||
|
||||
Canonical Structure findom_eqMixin T T' := EqMixin (@findom_eqP T T').
|
||||
Canonical Structure findom_eqType T T' := Eval hnf in EqType _ (findom_eqMixin T T').
|
||||
|
||||
Lemma leq_upd T (T':eqType) (f:FinDom T T') l a : l \notin (dom f) -> findom_leq f (updMap l a f).
|
||||
apply (findom_ind f) ; first by [].
|
||||
clear f. move => f n x A I C. rewrite indomUpdMap. case_eq (l == n) ; first by [].
|
||||
move => e. simpl. move => I'. specialize (C I'). rewrite updMapCom ; last by rewrite e.
|
||||
unfold findom_leq. apply (introT (@allP T _ _)). move => b. rewrite indomUpdMap.
|
||||
case_eq (b == n) => ee ; simpl.
|
||||
- move => _. rewrite (eqP ee). by do 2 rewrite updMap_simpl.
|
||||
- move => d. have N:(b != n) by case: (b == n) ee. do 2 rewrite (updMap_simpl2 _ _ N).
|
||||
have NN:(l != n) by case: (l == n) e. have F:(b == l) -> False. move => ex. rewrite (eqP ex) in d. by rewrite d in I'.
|
||||
have FF:(b != l). case: (b == l) F ; last by []. move => X. by case: (X is_true_true).
|
||||
by rewrite (updMap_simpl2 _ _ FF).
|
||||
Qed.
|
||||
|
||||
Lemma create_findomP (T:compType) (T':Type) (l:seq T) (f:T -> option T') : (uniq l) -> (sorted l) ->
|
||||
sorted (map (@fst _ _) (pmap (fun x : T => option_map [eta pair x] (f x)) l)) &&
|
||||
uniq (map (@fst _ _) (pmap (fun x : T => option_map [eta pair x] (f x)) l)).
|
||||
move => U S.
|
||||
have A:= (@pmap_uniq T _ _ id _ l U). rewrite map_pmap. rewrite -> A ; last by move => x ; simpl ; case (f x).
|
||||
rewrite andbT. clear A. clear U.
|
||||
elim: l S ; first by [].
|
||||
move => t s IH S. rewrite sorted_cons in S. specialize (IH (proj2 (andP S))).
|
||||
simpl. case (f t) ; last by apply IH.
|
||||
move => t'. simpl @oapp. rewrite sorted_cons. rewrite IH. rewrite andbT. clear IH t'.
|
||||
elim: s S ;first by []. move => x s IH A. rewrite sorted_cons in A. simpl in A.
|
||||
rewrite (proj2 (andP (proj1 (andP A)))) in IH. rewrite (proj2 (andP (proj2 (andP A)))) in IH. specialize (IH is_true_true).
|
||||
simpl. case (f x) ; last by apply IH. move => y. simpl. rewrite (proj1 (andP (proj1 (andP A)))). simpl.
|
||||
by apply IH.
|
||||
Qed.
|
||||
|
||||
Definition create_findom (T:compType) (T':Type) (l:seq T) (f:T -> option T') (U:uniq l) (S:sorted l) : FinDom T T' :=
|
||||
mk_findom (create_findomP f U S).
|
||||
|
||||
Lemma create_findom_simpl (T:compType) (T':Type) (l:seq T) (f:T -> option T') (U:uniq l) (S:sorted l) :
|
||||
findom_t (create_findom f U S) = (pmap (fun x => option_map (fun y => (x,y)) (f x)) l).
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition create_findomF (T:compType) (X T':Type) (l:FinDom T X) (f:T -> option T') : FinDom T T'.
|
||||
case:l f => l Pl f. by apply (create_findom f (proj2 (andP Pl)) (proj1 (andP Pl))).
|
||||
Defined.
|
||||
|
||||
Lemma dom_create_findom (T:compType) T' (f:T -> option T') x U S :
|
||||
dom (@create_findom _ _ x f U S) = filter [eta f] x.
|
||||
unfold dom.
|
||||
rewrite create_findom_simpl.
|
||||
rewrite pmap_filter. apply eq_filter => a. by case (f a).
|
||||
move => a. simpl. by case (f a).
|
||||
Qed.
|
||||
|
||||
Lemma dom_create_findomF (T:compType) T' T'' (f:T -> option T') (x:FinDom T T'') :
|
||||
dom (create_findomF x f) = filter [eta f] (dom x).
|
||||
case: x. simpl. move => x P. by rewrite dom_create_findom.
|
||||
Qed.
|
||||
|
||||
Lemma create_findomF_simpl (T:compType) T' T'' (f:T -> option T') (x:FinDom T T'') i :
|
||||
create_findomF x f i = if i \in dom x then f i else None.
|
||||
unfold findom_f. case: x => x P. simpl. unfold dom. simpl.
|
||||
elim: x P ; first by [].
|
||||
case => j x s IH. simpl @map. rewrite sorted_cons. simpl. move => P.
|
||||
rewrite (proj2 (andP (proj1 (andP P)))) in IH. rewrite (proj2 (andP (proj2 (andP P)))) in IH. specialize (IH is_true_true).
|
||||
case_eq (f j).
|
||||
- move => t fj. simpl. rewrite in_cons. case_eq (i == j) => IJ ; first by rewrite <- (eqP IJ) in fj ; rewrite fj.
|
||||
simpl. by apply IH.
|
||||
- move => fj. simpl. rewrite in_cons. rewrite IH. case_eq (i == j) => E ; last by [].
|
||||
simpl. rewrite <- (eqP E) in fj. rewrite fj. simpl. case_eq (i \in map (@fst _ _) s) => X ; by rewrite X.
|
||||
Qed.
|
||||
|
||||
|
||||
Section FinMap.
|
||||
|
||||
Variable T:compType.
|
||||
|
||||
Lemma Pemp T' (x:T) (P:x \in map (@fst _ T') nil) : False.
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Fixpoint indom_app_fi T' (s:seq (T * T')) x (P:x \in map (@fst _ _) s) :=
|
||||
match s as s0 return x \in map (@fst _ _) s0 -> T' with
|
||||
| nil => fun F => match Pemp F with end
|
||||
| (a,b)::r => fun P => (match (x == a) as b return b || (x \in map (@fst _ _) r) -> T'
|
||||
with true => fun _ => b
|
||||
| false => fun P => @indom_app_fi _ r x P end) P
|
||||
end P.
|
||||
|
||||
Definition indom_app T' (f:FinDom T T') x (P:x \in dom f) : T' := @indom_app_fi T' (findom_t f) x P.
|
||||
|
||||
Lemma indom_app_eq T' (f:FinDom T T') x (P:x \in dom f) : Some (indom_app P) = f x.
|
||||
case: f x P . move => s P x I. elim: s P I ; first by [].
|
||||
case => a b s IH. simpl @map. simpl @uniq. unfold dom. simpl @map. move => P. unfold in_mem. simpl.
|
||||
move => I. have X:= proj1 (andP P). rewrite sorted_cons in X.
|
||||
have Y:sorted (map (@fst _ _) s) && uniq (map (@fst _ _) s). rewrite (proj2 (andP X)). simpl.
|
||||
by rewrite (proj2 (andP (proj2 (andP P)))).
|
||||
specialize (IH Y). unfold dom in IH. simpl in IH.
|
||||
case_eq (x == a) => A.
|
||||
- move: I. unfold findom_f. simpl. move: A. clear. move => A. unfold indom_app. simpl. by rewrite A.
|
||||
- unfold findom_f. unfold indom_app. move: I. simpl. rewrite A. simpl. move => I. specialize (IH I).
|
||||
unfold indom_app in IH. simpl in IH. by rewrite IH.
|
||||
Qed.
|
||||
|
||||
Lemma mem_cons x y (s:seq T) : x \in s -> x \in (y::s).
|
||||
move => I. rewrite in_cons. rewrite I. by rewrite orbT.
|
||||
Qed.
|
||||
|
||||
Fixpoint indom_app_map T' P' (f:forall x, P' x -> T') (s:seq T) (I:forall x, x \in s -> P' x) : seq (T * T') :=
|
||||
match s as s0 return (forall x, x \in s0 -> P' x) -> seq (T * T') with
|
||||
| nil => fun _ => nil
|
||||
| cons x rs => fun P => (x,f x (P x (mem_head x rs))) :: (@indom_app_map _ P' f rs (fun y X => P _ (mem_cons x X)))
|
||||
end I.
|
||||
|
||||
Lemma list_fst_map T' P' (f:forall x, P' x -> T') (s:seq T) (I:forall x, x \in s -> P' x) :
|
||||
map (@fst _ _) (indom_app_map f I) = s.
|
||||
elim: s I ; first by [].
|
||||
move=> t s IH I. simpl. specialize (IH (fun (y : T) (X : y \in s) => I y (mem_cons t X))).
|
||||
f_equal. by apply IH.
|
||||
Qed.
|
||||
|
||||
Definition findom_map T' T'' (m:FinDom T T') (f:forall x, x \in dom m -> T'') : FinDom T T''.
|
||||
exists (@indom_app_map T'' (fun x => x \in dom m) f (dom m) (fun x X => X)).
|
||||
case: m f => s P f. rewrite list_fst_map. unfold dom. by apply P.
|
||||
Defined.
|
||||
|
||||
Lemma dom_findom_map T' T'' (m:FinDom T T') (f:forall x, x \in dom m -> T'') : dom m = dom (findom_map f).
|
||||
have A:=list_fst_map f (fun x => id). by rewrite <- A.
|
||||
Qed.
|
||||
|
||||
Lemma findom_fun_map T' P (f:forall x, P x -> T') (s:seq T) (I:forall x, x \in s -> P x) x (I':x \in s) :
|
||||
findom_fun (@indom_app_map _ P f s I) x = Some (f x (I _ I')).
|
||||
elim:s I x I' ; first by [].
|
||||
move => t s IH I x I'. simpl. case_eq (x == t) => E. move: I'. rewrite (eqP E). move => I'.
|
||||
by rewrite (eq_irrelevance (mem_head t s) I').
|
||||
specialize (IH (fun (y : T) (X : y \in s) => I y (mem_cons t X))). specialize (IH x).
|
||||
have I0:= I'. rewrite in_cons in I0. rewrite E in I0. simpl in I0. specialize (IH I0). rewrite IH.
|
||||
by rewrite (eq_irrelevance (mem_cons t I0) I').
|
||||
Qed.
|
||||
|
||||
Lemma findom_map_app T' T'' (m:FinDom T T') x (I:x \in dom m) (f:forall x, x \in dom m -> T'') :
|
||||
findom_map f x = Some (f x I).
|
||||
case: m x I f. move => s P. unfold findom_f. simpl. unfold dom. simpl.
|
||||
move => x I f. apply (findom_fun_map f (fun x => id) I).
|
||||
Qed.
|
||||
|
||||
End FinMap.
|
||||
|
173
papers/domains-2012/FinmapMetric.v
Normal file
173
papers/domains-2012/FinmapMetric.v
Normal file
|
@ -0,0 +1,173 @@
|
|||
(**********************************************************************************
|
||||
* Finmap.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
Require Export NSetoid MetricCore Finmap.
|
||||
|
||||
Section SET.
|
||||
Variable T:compType.
|
||||
Variable S:setoidType.
|
||||
|
||||
Lemma finmap_setoid_axiom : Setoid.axiom (fun f f' : FinDom T S => dom f =i dom f' /\ forall t, t \in dom f -> f t =-= f' t).
|
||||
split ; last split.
|
||||
- by move => f ; simpl ; split ; last by [].
|
||||
- move => x y z ; simpl => X Y. case: X => D X. case: Y => D' Y.
|
||||
split. move => a. rewrite (D a). by rewrite (D' a).
|
||||
move => a A. rewrite -> (X a A). rewrite D in A. by apply (Y _ A).
|
||||
- move => x y ; simpl => X. case: X => D X.
|
||||
split => a. by rewrite (D a). move => A. rewrite <- D in A. by rewrite -> (X a A).
|
||||
Qed.
|
||||
|
||||
Definition findom_setoidMixin := SetoidMixin finmap_setoid_axiom.
|
||||
Canonical Structure findom_setoidType := Eval hnf in SetoidType findom_setoidMixin.
|
||||
|
||||
Lemma indom_app_respect (f f': findom_setoidType) x (P:x \in dom f) (P':x \in dom f') : f =-= f' -> indom_app P =-= indom_app P'.
|
||||
move => e. have X:Some (indom_app P) =-= Some (indom_app P') ; last by apply X.
|
||||
do 2 rewrite indom_app_eq. case: f e P => s A. case: f' P' => s' A'. move => P e P'.
|
||||
unfold tset_eq in e. simpl in e. by apply (proj2 e).
|
||||
Qed.
|
||||
|
||||
Lemma respect_domeq (f f':findom_setoidType) : f =-= f' -> dom f =i dom f'.
|
||||
case:f => s P. case:f' => s' P'.
|
||||
unfold dom. simpl. unfold tset_eq. simpl. unfold dom. simpl. by move => [A _].
|
||||
Qed.
|
||||
|
||||
Lemma findom_sapp_respect (x:T) : setoid_respect (fun f:findom_setoidType => f x).
|
||||
move => f f' e. case_eq (x \in dom f) => D. by apply (proj2 e x D).
|
||||
rewrite findom_undef ; last by rewrite D. rewrite (proj1 e) in D.
|
||||
rewrite findom_undef ; last by rewrite D. by [].
|
||||
Qed.
|
||||
|
||||
Definition findom_sapp (x:T) : findom_setoidType =-> option_setoidType S :=
|
||||
Eval hnf in mk_fset (findom_sapp_respect x).
|
||||
|
||||
|
||||
End SET.
|
||||
|
||||
Section MET.
|
||||
Variable T:compType.
|
||||
Variable M:metricType.
|
||||
|
||||
Lemma findom_metric_axiom : Metric.axiom (fun n => (fun f f' : FinDom T M => match n with O => True | S _ =>
|
||||
dom f =i dom f' /\ forall i, i \in dom f -> i \in dom f' -> f i = n = f' i end)).
|
||||
move => n x y z. split ; last split ; last split ; last split ; clear.
|
||||
- split => X.
|
||||
+ split. specialize (X 1). simpl in X. by apply (proj1 X).
|
||||
move => t I. apply: (proj1 (Mrefl _ _)) => n. case: n ; first by []. move => n. specialize (X n.+1).
|
||||
simpl in X. apply (proj2 X). by apply I. rewrite (proj1 X) in I. by apply I.
|
||||
+ case ; first by []. move => n. split ; first by apply (proj1 X). move => i I I'.
|
||||
apply (proj2 (Mrefl _ _)). by apply (proj2 X _ I).
|
||||
- move => A. simpl. simpl in A. case: n A ; first by []. move => n [A B]. split ; first by move => a ; rewrite -> A.
|
||||
move => i I I'. specialize (B i I' I). by apply (Msym B).
|
||||
- move => A B. simpl in A,B. simpl. case: n A B ; first by [].
|
||||
move => n [A D] [A' D']. split ; first by move => a ; rewrite A; rewrite A'.
|
||||
move => i I I'. specialize (D i I). rewrite A in I. specialize (D I). specialize (D' i I I').
|
||||
by apply (Mrel_trans D D').
|
||||
- move => A. simpl. simpl in A. case: n A ; first by []. move => n [D A]. split ; first by [].
|
||||
move => i I I'. specialize (A i I I'). by apply Mmono.
|
||||
- by [].
|
||||
Qed.
|
||||
|
||||
Canonical Structure findom_metricMixin := MetricMixin findom_metric_axiom.
|
||||
Canonical Structure findom_metricType := Eval hnf in MetricType findom_metricMixin.
|
||||
|
||||
Lemma findom_f_nonexp (f f':findom_metricType) n x : f = n = f' -> f x = n = f' x.
|
||||
case: n x ; first by [].
|
||||
move => n x [E P]. case_eq (x \in dom f) => D.
|
||||
- have D':=D. rewrite E in D'. specialize (P x D D'). by apply P.
|
||||
rewrite findom_undef ; last by rewrite D.
|
||||
rewrite E in D. rewrite findom_undef ; last by rewrite D. by [].
|
||||
Qed.
|
||||
|
||||
Lemma findom_sapp_ne (x:T) : nonexpansive (fun f:findom_metricType => f x).
|
||||
move => n f f' e. by apply: findom_f_nonexp.
|
||||
Qed.
|
||||
|
||||
Definition findom_napp (x:T) : findom_metricType =-> option_metricType M :=
|
||||
Eval hnf in mk_fmet (findom_sapp_ne x).
|
||||
|
||||
End MET.
|
||||
|
||||
Section CMET.
|
||||
Variable T:compType.
|
||||
Variable M:cmetricType.
|
||||
|
||||
Lemma findom_chain_dom x n (c:cchain (findom_metricType T M)) : x \in dom (c 1) -> x \in dom (c n.+1).
|
||||
elim: n c ; first by [].
|
||||
move => n IH c X. specialize (IH (cutn 1 c)). simpl in IH. do 2 rewrite -> addSn, -> add0n in IH.
|
||||
apply IH. clear IH. have C:=cchain_cauchy c. specialize (C 1 1 2 (ltnSn _) (ltnW (ltnSn _))).
|
||||
by rewrite (proj1 C) in X.
|
||||
Qed.
|
||||
|
||||
Definition findom_chainx (c:cchain (findom_metricType T M)) x (P:x \in dom (c 1)) : cchain M.
|
||||
exists (fun n => @indom_app _ _ (c n.+1) _ (findom_chain_dom _ n _ P)).
|
||||
case ; first by [].
|
||||
move => n i j l l'.
|
||||
have X: Some (indom_app (findom_chain_dom x i c P)) = n.+1 = Some (indom_app (findom_chain_dom x j c P)) ; last by apply X.
|
||||
do 2 rewrite indom_app_eq. have a:= (@cchain_cauchy _ c n.+1 i.+1 j.+1 (ltnW l) (ltnW l')).
|
||||
by apply findom_f_nonexp.
|
||||
Defined.
|
||||
|
||||
Definition findom_lub (c:cchain (findom_metricType T M)) : findom_metricType T M :=
|
||||
@findom_map _ _ _ (c 1) (fun x X => umet_complete (@findom_chainx c x X)).
|
||||
|
||||
Lemma ff (P:T -> nat -> Prop) (A:forall x n m, n <= m -> P x n -> P x m) (s:seq T) :
|
||||
(forall x, x \in s -> exists m, P x m) -> exists m, forall x, x \in s -> P x m.
|
||||
elim: s ; first by move => X; exists 0.
|
||||
move => t s IH X. specialize (IH (fun x A => X x (mem_cons t A))). destruct IH as [m IH].
|
||||
have X':=X t (mem_head t _). destruct X' as [m' Pm']. exists (maxn m m'). move => x I.
|
||||
rewrite in_cons in I. case_eq (x == t) => E.
|
||||
- rewrite (eqP E). apply: (A _ _ _ _ Pm'). rewrite leq_maxr. rewrite leqnn. by rewrite orbT.
|
||||
- rewrite E in I. simpl in I. specialize (IH _ I). apply: (A _ _ _ _ IH). rewrite leq_maxr. by rewrite leqnn.
|
||||
Qed.
|
||||
|
||||
Lemma findom_lubP (c:cchain (findom_metricType T M)) : mconverge c (findom_lub c).
|
||||
unfold findom_lub. case ; first by exists 0. move => n.
|
||||
have A:exists m, forall x, x \in dom (c 1) -> forall (X: x \in dom (c 1)) i, m < i -> c i x = n.+1 = Some (umet_complete (findom_chainx _ _ X)).
|
||||
apply (@ff (fun x m => forall (X:x \in dom (c 1)) i, m < i -> c i x = n.+1 = Some (umet_complete (findom_chainx _ _ X)))).
|
||||
clear. move => x j i A Y X k L. apply Y. by apply (@ssrnat.leq_ltn_trans _ _ _ A L).
|
||||
move => x I. destruct (cumet_conv (findom_chainx _ _ I) n.+1) as [m P].
|
||||
exists m. move => X. case ; first by []. move => i L. specialize (P i L). simpl in P.
|
||||
rewrite <- (indom_app_eq (findom_chain_dom _ i _ I)).
|
||||
apply: (Mrel_trans P). apply umet_complete_extn. move=> j. simpl.
|
||||
have A:Some (indom_app (findom_chain_dom _ j _ I)) = n.+1 = Some (indom_app (findom_chain_dom _ j _ X)) ; last by apply A.
|
||||
rewrite (indom_app_eq (findom_chain_dom _ j _ I)). by rewrite (indom_app_eq (findom_chain_dom _ j _ X)).
|
||||
destruct A as [m A]. exists m.+1. case ; first by []. move => i L. split. rewrite <- dom_findom_map.
|
||||
have B:=cchain_cauchy c. specialize (B 1 i.+1 1 (ltn0Sn i) (ltn0Sn _)). apply (proj1 B).
|
||||
move => x I I'. specialize (A x).
|
||||
have J:x \in dom (c 1). have J:=cchain_cauchy c. specialize (J 1 i.+1 1 (ltn0Sn _) (ltn0Sn _)). have X:= (proj1 J).
|
||||
rewrite X in I. by [].
|
||||
specialize (A J J i.+1 L). apply (Mrel_trans A).
|
||||
by rewrite -> (findom_map_app J).
|
||||
Qed.
|
||||
|
||||
Canonical Structure findom_cmetricMixin := CMetricMixin findom_lubP.
|
||||
Canonical Structure findom_cmetricType := Eval hnf in CMetricType findom_cmetricMixin.
|
||||
|
||||
Lemma dom_findom_lub (c:cchain findom_cmetricType) : dom (umet_complete c) = dom (c 1).
|
||||
unfold umet_complete. simpl. unfold findom_lub.
|
||||
by rewrite <- dom_findom_map.
|
||||
Qed.
|
||||
|
||||
Lemma findom_lub_eq (c:cchain findom_cmetricType) x : umet_complete c x =-= umet_complete (liftc (findom_napp _ _ x) c).
|
||||
apply: (proj1 (Mrefl _ _)) => n.
|
||||
case: (cumet_conv c n). move => m X.
|
||||
case: (cumet_conv (liftc (findom_napp T M x) c) n). move => m' Y.
|
||||
specialize (X ((m+m')%N.+1)%N). specialize (X (leqW (leq_addr _ _))). specialize (Y ((m+m')%N.+1) (leqW (leq_addl _ _))).
|
||||
case: n X Y ; first by [].
|
||||
move => n X Y. case: X => D X. specialize (X x).
|
||||
- case_eq (x \in dom (umet_complete c)) => e.
|
||||
have e':=e. rewrite dom_findom_lub in e'. specialize (X (findom_chain_dom _ _ _ e') e).
|
||||
rewrite <- X. by rewrite <- Y.
|
||||
- rewrite findom_undef ; last by rewrite e. clear X. rewrite <- Y. clear Y. simpl.
|
||||
rewrite findom_undef ; first by []. rewrite D. by rewrite e.
|
||||
Qed.
|
||||
|
||||
End CMET.
|
||||
|
||||
|
||||
|
43
papers/domains-2012/KSOp.v
Normal file
43
papers/domains-2012/KSOp.v
Normal file
|
@ -0,0 +1,43 @@
|
|||
(**********************************************************************************
|
||||
* KSOp.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Operational semantics of the kitchen sink language *)
|
||||
|
||||
Require Import Finmap KSTm Fin.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
Import Tm.
|
||||
|
||||
Definition Heap := FinDom [compType of nat] (Value O).
|
||||
|
||||
Definition subSingle E (v:Value E) e := subExp (cons v (@idSub _)) e.
|
||||
|
||||
(*=EV *)
|
||||
Inductive EV : nat -> Exp 0 -> Heap -> Value O -> Heap -> Type :=
|
||||
| EvVAL h v : EV O (VAL v) h v h
|
||||
| EvFST h v0 v1 : EV O (FST (PAIR v0 v1)) h v0 h
|
||||
| EvSND h v0 v1 : EV O (SND (PAIR v0 v1)) h v1 h
|
||||
| EvOP h op n0 n1 : EV O (OP op (PAIR (INT n0) (INT n1))) h (INT (op n0 n1)) h
|
||||
| EvUNFOLD h v : EV 1 (UNFOLD (FOLD v)) h v h
|
||||
| EvREF (h:Heap) v (l:nat) : l \notin dom h -> EV 1 (REF v) h (LOC l) (updMap l v h)
|
||||
| EvBANG (h:Heap) v (l:nat) : h l = Some v -> EV 1 (BANG (LOC l)) h v h
|
||||
| EvASSIGN (h:Heap) v l : h l -> EV 1 (ASSIGN (LOC l) v) h UNIT (updMap l v h)
|
||||
| EvLET h n0 e0 v0 h0 n1 e1 v h1 : EV n0 e0 h v0 h0 ->
|
||||
EV n1 (subSingle v0 e1) h0 v h1 -> EV (n0 + n1) (LET e0 e1) h v h1
|
||||
| EvAPP h n e v0 v h0 : EV n (subSingle v e) h v0 h0 ->
|
||||
EV n (APP (LAM e) v) h v0 h0
|
||||
| EvTAPP h n e v0 h0 : EV n e h v0 h0 -> EV n (TAPP (TLAM e)) h v0 h0
|
||||
| EvCASEL h n v e0 e1 v0 h0 : EV n (subSingle v e0) h v0 h0 ->
|
||||
EV n (CASE (INL v) e0 e1) h v0 h0
|
||||
| EvCASER h n v e0 e1 v0 h0 : EV n (subSingle v e1) h v0 h0 ->
|
||||
EV n (CASE (INR v) e0 e1) h v0 h0.
|
||||
(*=End *)
|
||||
|
1000
papers/domains-2012/KSSem.v
Normal file
1000
papers/domains-2012/KSSem.v
Normal file
File diff suppressed because it is too large
Load diff
285
papers/domains-2012/KSTm.v
Normal file
285
papers/domains-2012/KSTm.v
Normal file
|
@ -0,0 +1,285 @@
|
|||
(**********************************************************************************
|
||||
* KSTy.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Kitchen sink language, well-scoped by construction *)
|
||||
|
||||
Require Export ssreflect ssrnat.
|
||||
Require Import Program.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
Require Import Fin.
|
||||
|
||||
Ltac Rewrites E :=
|
||||
(intros; simpl; try rewrite E;
|
||||
repeat (match goal with | [H : context[eq _ _] |- _] => rewrite H end);
|
||||
auto).
|
||||
|
||||
Module Tm.
|
||||
|
||||
Definition Env := nat.
|
||||
Definition Var := Fin.
|
||||
|
||||
(*=ValueExp *)
|
||||
Inductive Value E :=
|
||||
| VAR: Var E -> Value E
|
||||
| LOC: nat -> Value E
|
||||
| INT: nat -> Value E
|
||||
| TLAM: Exp E -> Value E
|
||||
| LAM: Exp E.+1 -> Value E
|
||||
| UNIT : Value E
|
||||
| PAIR: Value E -> Value E -> Value E
|
||||
| INL : Value E -> Value E
|
||||
| INR : Value E -> Value E
|
||||
| FOLD : Value E -> Value E
|
||||
with Exp E :=
|
||||
| VAL: Value E -> Exp E
|
||||
| FST: Value E -> Exp E
|
||||
| SND: Value E -> Exp E
|
||||
| OP: (nat -> nat -> nat) -> Value E -> Exp E
|
||||
| UNFOLD: Value E -> Exp E
|
||||
| REF: Value E -> Exp E
|
||||
| BANG: Value E -> Exp E
|
||||
| ASSIGN: Value E -> Value E -> Exp E
|
||||
| LET: Exp E -> Exp E.+1 -> Exp E
|
||||
| APP: Value E -> Value E -> Exp E
|
||||
| TAPP: Value E -> Exp E
|
||||
| CASE: Value E -> Exp E.+1 -> Exp E.+1 -> Exp E.
|
||||
(*=End *)
|
||||
|
||||
|
||||
Implicit Arguments INT [E].
|
||||
Implicit Arguments LOC [E].
|
||||
Implicit Arguments UNIT [E].
|
||||
|
||||
Scheme Value_ind2 := Induction for Value Sort Prop
|
||||
with Exp_ind2 := Induction for Exp Sort Prop.
|
||||
Combined Scheme ExpValue_ind from Value_ind2, Exp_ind2.
|
||||
|
||||
(*==========================================================================
|
||||
Variable-domain maps.
|
||||
By instantiating P with Var we get renamings.
|
||||
By instantiating P with Value we get substitutions.
|
||||
==========================================================================*)
|
||||
Module Map.
|
||||
|
||||
Section MAP.
|
||||
|
||||
Variable P : Env -> Type.
|
||||
Definition Map E E' := FMap E (P E').
|
||||
|
||||
(*==========================================================================
|
||||
Package of operations used with a Map
|
||||
vr maps a Var into Var or Value (so is either the identity or TVAR)
|
||||
vl maps a Var or Value to a Value (so is either TVAR or the identity)
|
||||
wk weakens a Var or Value (so is either SVAR or renaming through SVAR on a value)
|
||||
==========================================================================*)
|
||||
Record Ops :=
|
||||
{
|
||||
vr : forall E, Var E -> P E;
|
||||
vl : forall E, P E -> Value E;
|
||||
wk : forall E, P E -> P E.+1;
|
||||
wkvr : forall E (var : Var E), wk (vr var) = vr (FinS var);
|
||||
vlvr : forall E (var : Var E), vl (vr var) = VAR var
|
||||
}.
|
||||
Variable ops : Ops.
|
||||
|
||||
Definition lift E E' (m : Map E E') : Map E.+1 E'.+1 :=
|
||||
(fun var => match var in Fin E return Map E.-1 E' -> P E'.+1 with
|
||||
| FinZ _ => fun _ => vr ops (FinZ _)
|
||||
| FinS _ x => fun m => wk ops (m x)
|
||||
end m).
|
||||
|
||||
Definition shift E E' (m : Map E E') : Map E E'.+1 := fun var => wk ops (m var).
|
||||
Definition id E : Map E E := fun (var : Var E) => vr ops var.
|
||||
|
||||
Lemma shiftCons : forall E E' (m : Map E E') (x : P E'), shift (cons x m) = cons (wk ops x) (shift m).
|
||||
Proof. move => E E' m x. apply Extensionality. intros var. by dependent destruction var. Qed.
|
||||
|
||||
Lemma liftAsCons : forall E E' (m : Map E' E), lift m = cons (vr ops (FinZ _)) (shift m).
|
||||
Proof. move => E E' m. apply Extensionality. intros var. by dependent destruction var. Qed.
|
||||
|
||||
Fixpoint mapVal E E' (m : Map E E') (v : Value E) : Value E' :=
|
||||
match v with
|
||||
| VAR v => vl ops (m v)
|
||||
| INT i => INT i
|
||||
| LAM e => LAM (mapExp (lift m) e)
|
||||
| TLAM e => TLAM (mapExp m e)
|
||||
| LOC l => LOC l
|
||||
| UNIT => UNIT
|
||||
| PAIR v1 v2 => PAIR (mapVal m v1) (mapVal m v2)
|
||||
| INL v => INL (mapVal m v)
|
||||
| INR v => INR (mapVal m v)
|
||||
| FOLD v => FOLD (mapVal m v)
|
||||
end
|
||||
with mapExp E E' (m : Map E E') (e : Exp E) : Exp E' :=
|
||||
match e with
|
||||
| VAL v => VAL (mapVal m v)
|
||||
| APP v0 v1 => APP (mapVal m v0) (mapVal m v1)
|
||||
| LET e0 e1 => LET (mapExp m e0) (mapExp (lift m) e1)
|
||||
| FST v => FST (mapVal m v)
|
||||
| SND v => SND (mapVal m v)
|
||||
| UNFOLD v => UNFOLD (mapVal m v)
|
||||
| OP op v => OP op (mapVal m v)
|
||||
| REF v => REF (mapVal m v)
|
||||
| ASSIGN v1 v2 => ASSIGN (mapVal m v1) (mapVal m v2)
|
||||
| BANG v => BANG (mapVal m v)
|
||||
| TAPP v => TAPP (mapVal m v)
|
||||
| CASE v e1 e2 => CASE (mapVal m v) (mapExp (lift m) e1) (mapExp (lift m) e2)
|
||||
end.
|
||||
|
||||
Variable E E' : Env.
|
||||
Variable m : Map E E'.
|
||||
Lemma mapVAR : forall (var : Var _), mapVal m (VAR var) = vl ops (m var). by []. Qed.
|
||||
Lemma mapINT : forall n, mapVal m (INT n) = INT n. by []. Qed.
|
||||
Lemma mapLOC : forall n, mapVal m (LOC n) = LOC n. by []. Qed.
|
||||
Lemma mapLAM : forall (e : Exp _), mapVal m (LAM e) = LAM (mapExp (lift m) e). by []. Qed.
|
||||
Lemma mapOP : forall op v, mapExp m (OP op v) = OP op (mapVal m v). by []. Qed.
|
||||
Lemma mapVAL : forall (v : Value _), mapExp m (VAL v) = VAL (mapVal m v). by []. Qed.
|
||||
Lemma mapLET : forall (e1 : Exp _) (e2 : Exp _), mapExp m (LET e1 e2) = LET (mapExp m e1) (mapExp (lift m) e2). by []. Qed.
|
||||
Lemma mapAPP : forall (v1 : Value _) v2, mapExp m (APP v1 v2) = APP (mapVal m v1) (mapVal m v2). by []. Qed.
|
||||
Lemma mapPAIR : forall v1 v2, mapVal m (PAIR v1 v2) = PAIR (mapVal m v1) (mapVal m v2). by []. Qed.
|
||||
Lemma mapINL : forall v, mapVal m (INL v) = INL (mapVal m v). by []. Qed.
|
||||
Lemma mapINR : forall v, mapVal m (INR v) = INR (mapVal m v). by []. Qed.
|
||||
Lemma mapFOLD : forall v, mapVal m (FOLD v) = FOLD (mapVal m v). by []. Qed.
|
||||
Lemma mapTLAM : forall e, mapVal m (TLAM e) = TLAM (mapExp m e). by []. Qed.
|
||||
Lemma mapUNIT : mapVal m UNIT = UNIT. by []. Qed.
|
||||
Lemma mapFST : forall v, mapExp m (FST v) = FST (mapVal m v). by []. Qed.
|
||||
Lemma mapSND : forall v, mapExp m (SND v) = SND (mapVal m v). by []. Qed.
|
||||
Lemma mapBANG : forall v, mapExp m (BANG v) = BANG (mapVal m v). by []. Qed.
|
||||
Lemma mapREF : forall v, mapExp m (REF v) = REF (mapVal m v). by []. Qed.
|
||||
Lemma mapASSIGN : forall v1 v2, mapExp m (ASSIGN v1 v2) = ASSIGN (mapVal m v1) (mapVal m v2). by []. Qed.
|
||||
Lemma mapTAPP : forall v, mapExp m (TAPP v) = TAPP (mapVal m v). by []. Qed.
|
||||
Lemma mapUNFOLD : forall v, mapExp m (UNFOLD v) = UNFOLD (mapVal m v). by []. Qed.
|
||||
Lemma mapCASE : forall v e1 e2, mapExp m (CASE v e1 e2) = CASE (mapVal m v) (mapExp (lift m) e1) (mapExp (lift m) e2). by []. Qed.
|
||||
|
||||
Lemma liftId : lift (@id E) = @id E.+1.
|
||||
Proof. apply Extensionality. intros var. dependent destruction var; [by [] | by apply wkvr].
|
||||
Qed.
|
||||
|
||||
Lemma idAsCons : @id E.+1 = cons (vr ops (FinZ _)) (shift (@id E)).
|
||||
Proof. apply Extensionality. intros var. dependent destruction var; first by []. unfold id, shift. simpl. by rewrite wkvr. Qed.
|
||||
|
||||
End MAP.
|
||||
|
||||
Hint Rewrite mapVAR mapINT mapLAM mapOP mapVAL mapLET mapAPP mapPAIR mapINL mapINR mapFOLD mapTLAM mapUNIT mapFST mapSND mapBANG mapREF mapASSIGN mapTAPP mapUNFOLD mapCASE : mapHints.
|
||||
Implicit Arguments id [P].
|
||||
|
||||
Lemma applyId P (ops:Ops P) E :
|
||||
(forall (v : Value E), mapVal ops (id ops E) v = v)
|
||||
/\ (forall (e : Exp E), mapExp ops (id ops E) e = e).
|
||||
Proof. move: E ; apply ExpValue_ind; intros; autorewrite with mapHints; Rewrites liftId. by apply vlvr.
|
||||
Qed.
|
||||
|
||||
End Map.
|
||||
|
||||
(*==========================================================================
|
||||
Variable renamings: Map Var
|
||||
==========================================================================*)
|
||||
|
||||
Definition Ren := Map.Map Var.
|
||||
Definition RenOps : Map.Ops Var. refine (@Map.Build_Ops _ (fun _ v => v) VAR FinS _ _). by []. by []. Defined.
|
||||
|
||||
Definition renVal := Map.mapVal RenOps.
|
||||
Definition renExp := Map.mapExp RenOps.
|
||||
Definition liftRen := Map.lift RenOps.
|
||||
Definition shiftRen := Map.shift RenOps.
|
||||
Definition idRen := Map.id RenOps.
|
||||
|
||||
(*==========================================================================
|
||||
Composition of renaming
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeRen P E E' E'' (m : Map.Map P E' E'') (r : Ren E E') : Map.Map P E E'' := fun var => m (r var).
|
||||
|
||||
Lemma liftComposeRen : forall E E' E'' P ops (m:Map.Map P E' E'') (r:Ren E E'), Map.lift ops (composeRen m r) = composeRen (Map.lift ops m) (liftRen r).
|
||||
Proof. move => E E' E'' P ops m r. apply Extensionality. intros var. by dependent destruction var. Qed.
|
||||
|
||||
Lemma applyComposeRen E :
|
||||
(forall (v : Value E) E' E'' P ops (m:Map.Map P E' E'') (s : Ren E E'),
|
||||
Map.mapVal ops (composeRen m s) v = Map.mapVal ops m (renVal s v))
|
||||
/\ (forall (e : Exp E) E' E'' P ops (m:Map.Map P E' E'') (s : Ren E E'),
|
||||
Map.mapExp ops (composeRen m s) e = Map.mapExp ops m (renExp s e)).
|
||||
Proof.
|
||||
move: E ; apply ExpValue_ind; intros; autorewrite with mapHints; Rewrites liftComposeRen. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Substitution
|
||||
==========================================================================*)
|
||||
|
||||
Definition Sub := Map.Map Value.
|
||||
Definition SubOps : Map.Ops Value. refine (@Map.Build_Ops _ VAR (fun _ v => v) (fun E => renVal (fun v => FinS v)) _ _). by []. by []. Defined.
|
||||
|
||||
Definition subVal := Map.mapVal SubOps.
|
||||
Definition subExp := Map.mapExp SubOps.
|
||||
Definition shiftSub := Map.shift SubOps.
|
||||
Definition liftSub := Map.lift SubOps.
|
||||
Definition idSub := Map.id SubOps.
|
||||
|
||||
Ltac UnfoldRenSub := (unfold subVal; unfold subExp; unfold renVal; unfold renExp; unfold liftSub; unfold liftRen).
|
||||
Ltac FoldRenSub := (fold subVal; fold subExp; fold renVal; fold renExp; fold liftSub; fold liftRen).
|
||||
Ltac SimplMap := (UnfoldRenSub; autorewrite with mapHints; FoldRenSub).
|
||||
|
||||
(*==========================================================================
|
||||
Composition of substitution followed by renaming
|
||||
==========================================================================*)
|
||||
Definition composeRenSub E E' E'' (r : Ren E' E'') (s : Sub E E') : Sub E E'' := fun var => renVal r (s var).
|
||||
|
||||
Lemma liftComposeRenSub : forall E E' E'' (r:Ren E' E'') (s:Sub E E'), liftSub (composeRenSub r s) = composeRenSub (liftRen r) (liftSub s).
|
||||
intros. apply Extensionality. intros var. dependent destruction var; first by [].
|
||||
unfold composeRenSub, liftSub. simpl. unfold renVal at 1 3. by do 2 rewrite <- (proj1 (applyComposeRen _)).
|
||||
Qed.
|
||||
|
||||
Lemma applyComposeRenSub E :
|
||||
(forall (v:Value E) E' E'' (r:Ren E' E'') (s : Sub E E'),
|
||||
subVal (composeRenSub r s) v = renVal r (subVal s v))
|
||||
/\ (forall (e:Exp E) E' E'' (r:Ren E' E'') (s : Sub E E'),
|
||||
subExp (composeRenSub r s) e = renExp r (subExp s e)).
|
||||
Proof. move: E ; apply ExpValue_ind; intros; SimplMap; Rewrites liftComposeRenSub. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Composition of substitutions
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeSub E E' E'' (s' : Sub E' E'') (s : Sub E E') : Sub E E'' := fun var => subVal s' (s var).
|
||||
|
||||
Lemma liftComposeSub : forall E E' E'' (s' : Sub E' E'') (s : Sub E E'), liftSub (composeSub s' s) = composeSub (liftSub s') (liftSub s).
|
||||
intros. apply Extensionality. intros var. dependent destruction var; first by [].
|
||||
unfold composeSub. simpl. rewrite <- (proj1 (applyComposeRenSub _)). unfold composeRenSub, subVal. by rewrite <- (proj1 (applyComposeRen _)).
|
||||
Qed.
|
||||
|
||||
Lemma applyComposeSub E :
|
||||
(forall (v : Value E) E' E'' (s' : Sub E' E'') (s : Sub E E'),
|
||||
subVal (composeSub s' s) v = subVal s' (subVal s v))
|
||||
/\ (forall (e : Exp E) E' E'' (s' : Sub E' E'') (s : Sub E E'),
|
||||
subExp (composeSub s' s) e = subExp s' (subExp s e)).
|
||||
Proof. move: E ; apply ExpValue_ind; intros; SimplMap; Rewrites liftComposeSub. Qed.
|
||||
|
||||
Lemma composeCons : forall E E' E'' (s':Sub E' E'') (s:Sub E E') (v:Value _),
|
||||
composeSub (cons v s') (liftSub s) = cons v (composeSub s' s).
|
||||
intros. apply Extensionality. intros var. dependent destruction var; first by [].
|
||||
unfold composeSub. unfold subVal. simpl. rewrite <- (proj1 (applyComposeRen _)). unfold composeRen.
|
||||
simpl. replace ((fun (var0:Fin E') => s' var0)) with s' by (by apply Extensionality). by [].
|
||||
Qed.
|
||||
|
||||
Lemma composeSubIdLeft : forall E E' (s : Sub E E'), composeSub (@idSub _) s = s.
|
||||
Proof. intros. apply Extensionality. intros var. by apply (proj1 (Map.applyId _ _)). Qed.
|
||||
|
||||
Lemma composeSubIdRight : forall E E' (s:Sub E E'), composeSub s (@idSub _) = s.
|
||||
Proof. intros. by apply Extensionality. Qed.
|
||||
|
||||
Notation "[ x , .. , y ]" := (cons x .. (cons y (@Map.id _ SubOps _)) ..) : Sub_scope.
|
||||
Arguments Scope composeSub [_ _ _ Sub_scope Sub_scope].
|
||||
Arguments Scope subExp [_ _ Sub_scope].
|
||||
Arguments Scope subVal [_ _ Sub_scope].
|
||||
Delimit Scope Sub_scope with sub.
|
||||
|
||||
Lemma composeSingleSub : forall E E' (s:Sub E E') (v:Value _), composeSub [v] (liftSub s) = cons v s.
|
||||
Proof. intros. rewrite composeCons. by rewrite composeSubIdLeft. Qed.
|
||||
|
||||
End Tm.
|
229
papers/domains-2012/KSTy.v
Normal file
229
papers/domains-2012/KSTy.v
Normal file
|
@ -0,0 +1,229 @@
|
|||
(**********************************************************************************
|
||||
* KSTy.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Kitchen sink types *)
|
||||
|
||||
Require Export ssreflect ssrnat.
|
||||
Require Import Program.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
Require Import Fin.
|
||||
|
||||
Ltac Rewrites E :=
|
||||
(intros; simpl; try rewrite E;
|
||||
repeat (match goal with | [H : context[eq _ _] |- _] => rewrite H end);
|
||||
auto).
|
||||
|
||||
Module Ty.
|
||||
|
||||
Definition Env := nat.
|
||||
|
||||
(*=Ty *)
|
||||
Inductive Ty E :=
|
||||
| TVar: Fin E -> Ty E
|
||||
| Int: Ty E
|
||||
| Unit: Ty E
|
||||
| Product: Ty E -> Ty E -> Ty E
|
||||
| Sum: Ty E -> Ty E -> Ty E
|
||||
| Mu: Ty E.+1 -> Ty E
|
||||
| All: Ty E.+1 -> Ty E
|
||||
| Arrow: Ty E -> Ty E -> Ty E
|
||||
| Ref: Ty E -> Ty E.
|
||||
(*=End *)
|
||||
Implicit Arguments Unit [E].
|
||||
Implicit Arguments Int [E].
|
||||
|
||||
Scheme Ty_ind2 := Induction for Ty Sort Prop.
|
||||
|
||||
(*==========================================================================
|
||||
Variable-domain maps.
|
||||
By instantiating P with Var we get renamings.
|
||||
By instantiating P with Value we get substitutions.
|
||||
==========================================================================*)
|
||||
Module Map.
|
||||
Section MAP.
|
||||
|
||||
Variable P : Env -> Type.
|
||||
Definition Map E E' := FMap E (P E').
|
||||
|
||||
(*==========================================================================
|
||||
Package of operations used with a Map
|
||||
vr maps a Var into Var or Value (so is either the identity or TVAR)
|
||||
vl maps a Var or Value to a Value (so is either TVAR or the identity)
|
||||
wk weakens a Var or Value (so is either SVAR or renaming through SVAR on a value)
|
||||
==========================================================================*)
|
||||
Record Ops :=
|
||||
{
|
||||
vr : forall E, Fin E -> P E;
|
||||
vl : forall E, P E -> Ty E;
|
||||
wk : forall E, P E -> P E.+1;
|
||||
wkvr : forall E (var : Fin E), wk (vr var) = vr (FinS var);
|
||||
vlvr : forall E (var : Fin E), vl (vr var) = TVar var
|
||||
}.
|
||||
Variable ops : Ops.
|
||||
|
||||
|
||||
Definition lift E E' (m : Map E E') : Map E.+1 E'.+1 :=
|
||||
(fun var => match var in Fin E return Map E.-1 E' -> P E'.+1 with
|
||||
| FinZ _ => fun _ => vr ops (FinZ _)
|
||||
| FinS _ x => fun m => wk ops (m x)
|
||||
end m).
|
||||
|
||||
Definition shift E E' (m : Map E E') : Map E E'.+1 := fun var => wk ops (m var).
|
||||
Definition id E : Map E E := fun (var : Fin E) => vr ops var.
|
||||
|
||||
Lemma shiftCons : forall E E' (m : Map E E') (x : P E'), shift (cons x m) = cons (wk ops x) (shift m).
|
||||
Proof. move => E E' m x. apply Extensionality. by dependent destruction i. Qed.
|
||||
|
||||
Lemma liftAsCons : forall E E' (m : Map E' E), lift m = cons (vr ops (FinZ _)) (shift m).
|
||||
Proof. move => E E' m. apply Extensionality. by dependent destruction i. Qed.
|
||||
|
||||
Fixpoint mapTy E E' (m : Map E E') (t : Ty E) : Ty E' :=
|
||||
match t with
|
||||
| TVar v => vl ops (m v)
|
||||
| Int => Int
|
||||
| Unit => Unit
|
||||
| Product t1 t2 => Product (mapTy m t1) (mapTy m t2)
|
||||
| Sum t1 t2 => Sum (mapTy m t1) (mapTy m t2)
|
||||
| Mu t => Mu (mapTy (lift m) t)
|
||||
| All t => All (mapTy (lift m) t)
|
||||
| Arrow t1 t2 => Arrow (mapTy m t1) (mapTy m t2)
|
||||
| Ref t => Ref (mapTy m t)
|
||||
end.
|
||||
|
||||
Variable E E' : Env.
|
||||
Variable m : Map E E'.
|
||||
Lemma mapTVar : forall (var : Fin _), mapTy m (TVar var) = vl ops (m var). by []. Qed.
|
||||
Lemma mapInt : mapTy m Int = Int. by []. Qed.
|
||||
Lemma mapUnit : mapTy m Unit = Unit. by []. Qed.
|
||||
Lemma mapMu : forall t, mapTy m (Mu t) = Mu (mapTy (lift m) t). by []. Qed.
|
||||
Lemma mapAll : forall t, mapTy m (All t) = All (mapTy (lift m) t). by []. Qed.
|
||||
Lemma mapProduct : forall t1 t2, mapTy m (Product t1 t2) = Product (mapTy m t1) (mapTy m t2). by []. Qed.
|
||||
Lemma mapSum : forall t1 t2, mapTy m (Sum t1 t2) = Sum (mapTy m t1) (mapTy m t2). by []. Qed.
|
||||
Lemma mapArrow : forall t1 t2, mapTy m (Arrow t1 t2) = Arrow (mapTy m t1) (mapTy m t2). by []. Qed.
|
||||
Lemma mapRef : forall t, mapTy m (Ref t) = Ref (mapTy m t). by []. Qed.
|
||||
|
||||
Lemma liftId : lift (@id E) = @id E.+1.
|
||||
Proof. apply Extensionality. dependent destruction i; [by [] | by apply wkvr].
|
||||
Qed.
|
||||
|
||||
Lemma idAsCons : @id E.+1 = cons (vr ops (FinZ _)) (shift (@id E)).
|
||||
Proof. apply Extensionality. dependent destruction i; first by []. unfold id, shift. simpl. by rewrite wkvr. Qed.
|
||||
|
||||
End MAP.
|
||||
|
||||
Hint Rewrite mapTVar mapInt mapUnit mapMu mapAll mapProduct mapSum mapArrow mapRef : mapHints.
|
||||
Implicit Arguments id [P].
|
||||
|
||||
Lemma applyId P (ops:Ops P) E : (forall (t : Ty E), mapTy ops (id ops E) t = t).
|
||||
Proof. induction t; intros; autorewrite with mapHints; Rewrites liftId. by apply vlvr. Qed.
|
||||
|
||||
End Map.
|
||||
|
||||
(*==========================================================================
|
||||
Variable renamings: Map Var
|
||||
==========================================================================*)
|
||||
|
||||
Definition Ren := Map.Map Fin.
|
||||
Definition RenOps : Map.Ops Fin. refine (@Map.Build_Ops _ (fun _ v => v) TVar FinS _ _). by []. by []. Defined.
|
||||
|
||||
Definition renTy := Map.mapTy RenOps.
|
||||
Definition liftRen := Map.lift RenOps.
|
||||
Definition shiftRen := Map.shift RenOps.
|
||||
Definition idRen := Map.id RenOps.
|
||||
|
||||
(*==========================================================================
|
||||
Composition of renaming
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeRen P E E' E'' (m : Map.Map P E' E'') (r : Ren E E') : Map.Map P E E'' := fun var => m (r var).
|
||||
|
||||
Lemma liftComposeRen : forall E E' E'' P ops (m:Map.Map P E' E'') (r:Ren E E'), Map.lift ops (composeRen m r) = composeRen (Map.lift ops m) (liftRen r).
|
||||
Proof. move => E E' E'' P ops m r. apply Extensionality. by dependent destruction i. Qed.
|
||||
|
||||
Lemma applyComposeRen E :
|
||||
(forall (t : Ty E) E' E'' P ops (m:Map.Map P E' E'') (s : Ren E E'),
|
||||
Map.mapTy ops (composeRen m s) t = Map.mapTy ops m (renTy s t)).
|
||||
Proof. induction t; intros; autorewrite with mapHints; Rewrites liftComposeRen. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Substitution
|
||||
==========================================================================*)
|
||||
|
||||
Definition Sub := Map.Map Ty.
|
||||
Definition SubOps : Map.Ops Ty. refine (@Map.Build_Ops _ TVar (fun _ v => v) (fun E => renTy (fun v => FinS v)) _ _). by []. by []. Defined.
|
||||
|
||||
Definition subTy := Map.mapTy SubOps.
|
||||
Definition shiftSub := Map.shift SubOps.
|
||||
Definition liftSub := Map.lift SubOps.
|
||||
Definition idSub := Map.id SubOps.
|
||||
|
||||
Ltac UnfoldRenSub := (unfold subTy; unfold renTy; unfold liftSub; unfold liftRen).
|
||||
Ltac FoldRenSub := (fold subTy; fold renTy; fold liftSub; fold liftRen).
|
||||
Ltac SimplMap := (UnfoldRenSub; autorewrite with mapHints; FoldRenSub).
|
||||
|
||||
(*==========================================================================
|
||||
Composition of substitution followed by renaming
|
||||
==========================================================================*)
|
||||
Definition composeRenSub E E' E'' (r : Ren E' E'') (s : Sub E E') : Sub E E'' := fun var => renTy r (s var).
|
||||
|
||||
Lemma liftComposeRenSub : forall E E' E'' (r:Ren E' E'') (s:Sub E E'), liftSub (composeRenSub r s) = composeRenSub (liftRen r) (liftSub s).
|
||||
intros. apply Extensionality. dependent destruction i; first by [].
|
||||
unfold composeRenSub, liftSub. simpl. unfold renTy at 1 3. by do 2 rewrite <- applyComposeRen.
|
||||
Qed.
|
||||
|
||||
Lemma applyComposeRenSub E :
|
||||
(forall (t:Ty E) E' E'' (r:Ren E' E'') (s : Sub E E'),
|
||||
subTy (composeRenSub r s) t = renTy r (subTy s t)).
|
||||
Proof. induction t; intros; SimplMap; Rewrites liftComposeRenSub. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Composition of substitutions
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeSub E E' E'' (s' : Sub E' E'') (s : Sub E E') : Sub E E'' := fun var => subTy s' (s var).
|
||||
|
||||
Lemma liftComposeSub : forall E E' E'' (s' : Sub E' E'') (s : Sub E E'), liftSub (composeSub s' s) = composeSub (liftSub s') (liftSub s).
|
||||
intros. apply Extensionality. dependent destruction i; first by [].
|
||||
unfold composeSub. simpl. rewrite <- applyComposeRenSub. unfold composeRenSub, subTy. by rewrite <- applyComposeRen.
|
||||
Qed.
|
||||
|
||||
Lemma applyComposeSub E :
|
||||
(forall (t : Ty E) E' E'' (s' : Sub E' E'') (s : Sub E E'),
|
||||
subTy (composeSub s' s) t = subTy s' (subTy s t)).
|
||||
Proof. induction t; intros; SimplMap; Rewrites liftComposeSub. Qed.
|
||||
|
||||
Lemma composeCons : forall E E' E'' (s':Sub E' E'') (s:Sub E E') (v:Ty _),
|
||||
composeSub (cons v s') (liftSub s) = cons v (composeSub s' s).
|
||||
intros. apply Extensionality. dependent destruction i; first by [].
|
||||
unfold composeSub. unfold subTy. simpl. rewrite <- applyComposeRen. unfold composeRen.
|
||||
simpl. replace ((fun (var0:Fin E') => s' var0)) with s' by (by apply Extensionality). by [].
|
||||
Qed.
|
||||
|
||||
Lemma composeSubIdLeft : forall E E' (s : Sub E E'), composeSub (@idSub _) s = s.
|
||||
Proof. intros. apply Extensionality. intros var. by apply Map.applyId. Qed.
|
||||
|
||||
Lemma composeSubIdRight : forall E E' (s:Sub E E'), composeSub s (@idSub _) = s.
|
||||
Proof. intros. by apply Extensionality. Qed.
|
||||
|
||||
Notation "[ x , .. , y ]" := (cons x .. (cons y (@Map.id _ SubOps _)) ..) : Sub_scope.
|
||||
Arguments Scope composeSub [_ _ _ Sub_scope Sub_scope].
|
||||
Arguments Scope subTy [_ _ Sub_scope].
|
||||
Delimit Scope Sub_scope with sub.
|
||||
|
||||
Lemma composeSingleSub : forall E E' (s:Sub E E') (t:Ty _), composeSub [t] (liftSub s) = cons t s.
|
||||
Proof. intros. rewrite composeCons. by rewrite composeSubIdLeft. Qed.
|
||||
|
||||
End Ty.
|
||||
|
||||
Notation "a --> b" := (Ty.Arrow a b) (at level 55, right associativity).
|
||||
Notation "a ** b" := (Ty.Product a b) (at level 55).
|
||||
|
||||
|
||||
|
70
papers/domains-2012/KSTyping.v
Normal file
70
papers/domains-2012/KSTyping.v
Normal file
|
@ -0,0 +1,70 @@
|
|||
(**********************************************************************************
|
||||
* KSTyping.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Typing relation for kitchen sink language *)
|
||||
|
||||
Require Export ssreflect ssrnat. Require Import KSTy. Require Import KSTm.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
Require Import Fin.
|
||||
|
||||
Definition StoreType E := nat -> Ty.Ty E.
|
||||
|
||||
Definition TEnv E n := FMap n (Ty.Ty E).
|
||||
Definition shiftTy E : Ty.Ty E -> Ty.Ty E.+1 := Ty.renTy (fun v => FinS v).
|
||||
Definition subOneTy E (t' : Ty.Ty E) t := Ty.subTy (cons t' (@Ty.idSub _)) t.
|
||||
Definition unfoldTy E (t : Ty.Ty E.+1) := subOneTy (Ty.Mu t) t.
|
||||
|
||||
Require Import Program.
|
||||
|
||||
Definition emptyTEnv : TEnv O O.
|
||||
move => var. dependent destruction var.
|
||||
Defined.
|
||||
|
||||
Import Tm.
|
||||
Inductive VTy E (se:StoreType E) : forall n, TEnv E n -> Tm.Value n -> Ty.Ty E -> Type :=
|
||||
| TvVAR: forall n (env : TEnv E n) v, VTy se env (VAR v) (env v)
|
||||
| TvLOC: forall n (env : TEnv E n) l, VTy se env (LOC l) (Ty.Ref (se l))
|
||||
| TvINT: forall n (env : TEnv E n) i , VTy se env (INT i) Ty.Int
|
||||
| TvUNIT: forall n (env : TEnv E n), VTy se env UNIT Ty.Unit
|
||||
| TvLAM: forall n (env : TEnv E n) t1 t2 e, ETy se (cons t1 env) e t2 -> VTy se env (LAM e) (t1 --> t2)
|
||||
| TvTLAM: forall n (env : TEnv E n) t e, ETy (E:=E.+1) (fun l => shiftTy (se l)) (fun v => shiftTy (env v)) e t -> VTy se env (TLAM e) (Ty.All t)
|
||||
| TvPAIR: forall n (env : TEnv E n) t1 t2 e1 e2, VTy se env e1 t1 -> VTy se env e2 t2 -> VTy se env (PAIR e1 e2) (t1 ** t2)
|
||||
| TvINL: forall n (env:TEnv E n) v t1 t2, VTy se env v t1 -> VTy se env (INL v) (Ty.Sum t1 t2)
|
||||
| TvINR: forall n (env:TEnv E n) v t1 t2, VTy se env v t2 -> VTy se env (INR v) (Ty.Sum t1 t2)
|
||||
| TvFOLD: forall n (env:TEnv E n) v t, VTy se env v (unfoldTy t) -> VTy se env (FOLD v) (Ty.Mu t)
|
||||
|
||||
with ETy E (se:StoreType E) : forall n, TEnv E n -> Tm.Exp n -> Ty.Ty E -> Type :=
|
||||
| TeVAL: forall n (env:TEnv E n) v t, VTy se env v t -> ETy se env (VAL v) t
|
||||
| TeLET: forall n (env:TEnv E n) e1 e2 t1 t2, ETy se env e1 t1 -> ETy se (cons t1 env) e2 t2 -> ETy se env (LET e1 e2) t2
|
||||
| TvFST: forall n (env:TEnv E n) v t1 t2, VTy se env v (t1**t2) -> ETy se env (FST v) t1
|
||||
| TvSND: forall n (env:TEnv E n) v t1 t2, VTy se env v (t1**t2) -> ETy se env (SND v) t2
|
||||
| TvOP: forall n (env:TEnv E n) op v, VTy se env v (Ty.Int**Ty.Int) -> ETy se env (OP op v) Ty.Int
|
||||
| TvUNFOLD: forall n (env:TEnv E n) t v, VTy se env v (Ty.Mu t) -> ETy se env (UNFOLD v) (unfoldTy t)
|
||||
| TvREF: forall n (env:TEnv E n) v t, VTy se env v t -> ETy se env (REF v) (Ty.Ref t)
|
||||
| TvBANG: forall n (env:TEnv E n) v t, VTy se env v (Ty.Ref t) -> ETy se env (BANG v) t
|
||||
| TvASSIGN: forall n (env:TEnv E n) v1 v2 t, VTy se env v1 (Ty.Ref t) -> VTy se env v2 t -> ETy se env (ASSIGN v1 v2) Ty.Unit
|
||||
| TeAPP: forall n (env:TEnv E n) t1 t2 v1 v2, VTy se env v1 (t1 --> t2) -> VTy se env v2 t1 -> ETy se env (APP v1 v2) t2
|
||||
| TeTAPP: forall n (env:TEnv E n) t v t', VTy se env v (Ty.All t) -> ETy se env (TAPP v) (subOneTy t' t)
|
||||
| TeCASE: forall n (env:TEnv E n) t1 t2 v e1 e2 t, VTy se env v (Ty.Sum t1 t2) -> ETy se (cons t1 env) e1 t -> ETy se (cons t2 env) e2 t -> ETy se env (CASE v e1 e2) t
|
||||
.
|
||||
|
||||
Hint Resolve TvINT TeAPP TeCASE TeLET TeVAL TvVAR TvTLAM TLAM TvOP TvINL TvINR TvPAIR TvSND TvFST TvASSIGN TvBANG TvREF.
|
||||
|
||||
Scheme VTy_rec2 := Induction for VTy Sort Set
|
||||
with ETy_rec2 := Induction for ETy Sort Set.
|
||||
|
||||
Scheme VTy_rect2 := Induction for VTy Sort Type
|
||||
with ETy_rect2 := Induction for ETy Sort Type.
|
||||
|
||||
Scheme VTy_ind2 := Induction for VTy Sort Prop
|
||||
with ETy_ind2 := Induction for ETy Sort Prop.
|
||||
|
||||
Combined Scheme Typing_ind from VTy_ind2, ETy_ind2.
|
||||
|
532
papers/domains-2012/KnasterTarski.v
Executable file
532
papers/domains-2012/KnasterTarski.v
Executable file
|
@ -0,0 +1,532 @@
|
|||
(**********************************************************************************
|
||||
* KnasterTarski.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Complete Lattices *)
|
||||
|
||||
Require Export PredomCore.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
(** printing <= $\sqsubseteq$ *)
|
||||
(** printing sup $\bigwedge$ *)
|
||||
(** printing inf $\bigvee$ *)
|
||||
(** printing == $\equiv$ *)
|
||||
|
||||
(** * Complete Lattice
|
||||
A semi-lattice [(L, \/ )] is a partial order [L = (T, <= )] and a binary operation inf [\/ : L -> L -> L],
|
||||
such that [forall x y, x \/ y <= x], [forall x y, x \/ y == y \/ x] and [forall z x y, z <= x /\ z <= y -> z <= x \/ y].
|
||||
|
||||
A complete lattice is a semi lattice where the inf operation has been generalized to all subsets of [L].
|
||||
*)
|
||||
|
||||
Open Scope C_scope.
|
||||
|
||||
Module CLattice.
|
||||
|
||||
Section Mixin.
|
||||
|
||||
Variable T:ordType.
|
||||
|
||||
Definition axiom (inf : (T -> Prop) -> T) :=
|
||||
forall P:T -> Prop, (forall t, (P t -> inf P <= t) /\ ((forall x, P x -> t <= x) -> t <= inf P)).
|
||||
|
||||
Record mixin_of : Type := Mixin
|
||||
{ inf : (T -> Prop) -> T;
|
||||
_ : axiom inf
|
||||
}.
|
||||
|
||||
End Mixin.
|
||||
|
||||
Section ClassDef.
|
||||
Record class_of (T:Type) : Type :=
|
||||
Class { base : PreOrd.class_of T ;
|
||||
ext : mixin_of (PreOrd.Pack base T) }.
|
||||
Local Coercion base : class_of >-> PreOrd.class_of.
|
||||
Local Coercion ext : class_of >-> mixin_of.
|
||||
|
||||
Structure type : Type := Pack { sort : Type; _ : class_of sort ; _ : Type}.
|
||||
Local Coercion sort : type >-> Sortclass.
|
||||
|
||||
Definition class cT := let: Pack _ c _ := cT return class_of cT in c.
|
||||
Definition unpack K (k : forall T (c : class_of T), K T c) cT :=
|
||||
let: Pack T c _ := cT return K _ (class cT) in k _ c.
|
||||
Definition repack cT : _ -> Type -> type := let k T c p := p c in unpack k cT.
|
||||
|
||||
Definition pack := let k T c m := Pack (@Class T c m) T in PreOrd.unpack k.
|
||||
|
||||
Definition ordType cT := PreOrd.Pack (class cT) cT.
|
||||
Definition setoidType cT := Setoid.Pack (class cT) cT.
|
||||
End ClassDef.
|
||||
Module Import Exports.
|
||||
Coercion base : class_of >-> PreOrd.class_of.
|
||||
Coercion ext : class_of >-> mixin_of.
|
||||
Coercion sort : type >-> Sortclass.
|
||||
Coercion ordType : type >-> PreOrd.type.
|
||||
Coercion setoidType : type >-> Setoid.type.
|
||||
Notation clatType := type.
|
||||
Notation ClatMixin := Mixin.
|
||||
Notation ClatType := pack.
|
||||
|
||||
Canonical Structure ordType.
|
||||
Canonical Structure setoidType.
|
||||
End Exports.
|
||||
End CLattice.
|
||||
Export CLattice.Exports.
|
||||
|
||||
Definition inf (cT:clatType) : (cT -> Prop) -> cT := CLattice.inf (CLattice.class cT).
|
||||
Implicit Arguments inf [cT].
|
||||
|
||||
Lemma Pinf (T:clatType) : forall P : T -> Prop, forall t, P t -> inf P <= t.
|
||||
unfold inf. case:T => A. case => O. simpl. case. simpl.
|
||||
move => I X a P t x. by apply (proj1 (X P t) x).
|
||||
Qed.
|
||||
|
||||
Lemma PinfL (T:clatType) : forall P : T -> Prop, forall t, (forall x, P x -> t <= x) -> t <= inf P.
|
||||
unfold inf. case:T => A. case => O. simpl. case. simpl.
|
||||
move => I X T P t x. by apply (proj2 (X P t) x).
|
||||
Qed.
|
||||
|
||||
Lemma inf_ext (T:clatType) (P P' : T -> Prop) : (forall x, P x <-> P' x) -> inf P =-= inf P'.
|
||||
move => C. split.
|
||||
- apply: PinfL. move => x Px'. apply Pinf. by apply (proj2 (C x) Px').
|
||||
- apply: PinfL. move => x Px'. apply Pinf. by apply (proj1 (C x) Px').
|
||||
Qed.
|
||||
|
||||
|
||||
Add Parametric Relation (O:clatType) : O (@tset_eq O : O -> O -> Prop)
|
||||
reflexivity proved by (@Oeq_refl O) symmetry proved by (@Oeq_sym O)
|
||||
transitivity proved by (@Oeq_trans O) as Oeq_ClatRelation.
|
||||
|
||||
Add Parametric Relation (O:clatType) : O (@Ole O)
|
||||
reflexivity proved by (@Ole_refl O)
|
||||
transitivity proved by (@Ole_trans O) as Ole_ClatRelation.
|
||||
|
||||
Definition imageSet (T0 T1:clatType) (f:ordCatType T0 T1) : (T0 -> Prop) -> T1 -> Prop :=
|
||||
fun S x => exists y, S y /\ f y = x.
|
||||
|
||||
Definition limitpres (T0 T1:clatType) (f: ordCatType T0 T1) :=
|
||||
forall S, f (inf S) =-= inf (imageSet f S).
|
||||
|
||||
Module FLimit.
|
||||
|
||||
Section flimit.
|
||||
Variable O1 O2 : clatType.
|
||||
|
||||
Record mixin_of (f:fmono O1 O2) := Mixin { cont :> limitpres f }.
|
||||
|
||||
Record class_of (f : O1 -> O2) := Class { base : FMon.mixin_of f; ext : mixin_of (FMon.Pack base f) }.
|
||||
Local Coercion base : class_of >-> FMon.mixin_of.
|
||||
Local Coercion ext : class_of >-> mixin_of.
|
||||
|
||||
Structure type : Type := Pack {sort : O1 -> O2; _ : class_of sort; _ : O1 -> O2}.
|
||||
Local Coercion sort : type >-> Funclass.
|
||||
Definition class cT := let: Pack _ c _ := cT return class_of cT in c.
|
||||
Definition unpack K (k : forall T (c : class_of T), K T c) cT :=
|
||||
let: Pack T c _ := cT return K _ (class cT) in k _ c.
|
||||
Definition repack cT : _ -> Type -> type := let k T c p := p c in unpack k cT.
|
||||
Definition pack := let k T c m := Pack (@Class T c m) T in FMon.unpack k.
|
||||
|
||||
Definition fmono f : fmono O1 O2 := FMon.Pack (class f) f.
|
||||
End flimit.
|
||||
Module Import Exports.
|
||||
Coercion base : class_of >-> FMon.mixin_of.
|
||||
Coercion ext : class_of >-> mixin_of.
|
||||
Coercion sort : type >-> Funclass.
|
||||
Coercion fmono : type >-> FMon.type.
|
||||
Notation flimit := type.
|
||||
Notation flimitMixin := Mixin.
|
||||
Notation flimitType := pack.
|
||||
Canonical Structure fmono.
|
||||
End Exports.
|
||||
End FLimit.
|
||||
Export FLimit.Exports.
|
||||
|
||||
Lemma flimitpres (T0 T1:clatType) (f:flimit T0 T1) : limitpres f.
|
||||
case:f => f. case ; case => m. case => l s. by apply l.
|
||||
Qed.
|
||||
|
||||
Lemma flimitp_ord_axiom (D1 D2 :clatType) : PreOrd.axiom (fun (f g:flimit D1 D2) => (f:fmono D1 D2) <= g).
|
||||
move => f ; split ; first by [].
|
||||
move => g h ; by apply Ole_trans.
|
||||
Qed.
|
||||
|
||||
Canonical Structure flimitp_ordMixin (D1 D2 :clatType) := OrdMixin (@flimitp_ord_axiom D1 D2).
|
||||
Canonical Structure flimitp_ordType (D1 D2 :clatType) := Eval hnf in OrdType (flimitp_ordMixin D1 D2).
|
||||
|
||||
Lemma clatMorphSetoidAxiom (O0 O1:clatType) : @Setoid.axiom (flimit O0 O1) (fun x y => (x:fmono O0 O1) =-= y).
|
||||
split ; first by move => x ; apply: tset_refl. split.
|
||||
- by apply: Oeq_trans.
|
||||
- by apply: Oeq_sym.
|
||||
Qed.
|
||||
|
||||
Canonical Structure clatMorphSetoidMixin (T0 T1 : clatType) := SetoidMixin (@clatMorphSetoidAxiom T0 T1).
|
||||
Canonical Structure clatMorphSetoidType (T0 T1 : clatType) := Eval hnf in SetoidType (clatMorphSetoidMixin T0 T1).
|
||||
|
||||
(* flimit <= *)
|
||||
Definition flimit_less (A B: clatType): relation (flimit A B) := (@Ole _).
|
||||
|
||||
Definition flimit_less_preorder (A B: clatType): PreOrder (@flimit_less A B).
|
||||
split ;first by move => x.
|
||||
move => x y z. by apply Ole_trans.
|
||||
Defined.
|
||||
|
||||
Existing Instance flimit_less_preorder.
|
||||
|
||||
Add Parametric Morphism (A B : clatType) :
|
||||
(@FLimit.sort A B) with signature (@flimit_less A B ++> @Ole A ++> @Ole B)
|
||||
as flimit_le_compat.
|
||||
move => x y l x' y' l'. apply Ole_trans with (y:=x y') ; first by apply: fmonotonic.
|
||||
by apply l.
|
||||
Qed.
|
||||
|
||||
(* fcont == *)
|
||||
Definition flimit_equal (A B: clatType): relation (flimit A B) := (@tset_eq _).
|
||||
|
||||
Definition flimit_equal_equivalence (A B: clatType): Equivalence (@flimit_equal A B).
|
||||
split.
|
||||
- move => x. by apply: Oeq_refl.
|
||||
- move => x y. by apply: Oeq_sym.
|
||||
- move => x y z. by apply: Oeq_trans.
|
||||
Defined.
|
||||
|
||||
Existing Instance flimit_equal_equivalence.
|
||||
|
||||
Add Parametric Morphism (A B : clatType) :
|
||||
(@FLimit.sort A B) with signature (@flimit_equal A B ==> @tset_eq A ==> @tset_eq B)
|
||||
as flimit_eq_compat.
|
||||
move => x y l x' y' l'. apply Oeq_trans with (y:=x y') ; first by apply: fmon_stable.
|
||||
by apply (fmon_eq_elim l y').
|
||||
Qed.
|
||||
|
||||
Lemma fcomplp (D0 D1 D2 : clatType) (f : flimit D1 D2) (g: flimit D0 D1) :
|
||||
limitpres (ocomp f g).
|
||||
move => S. simpl.
|
||||
have gg:=flimitpres g S. rewrite -> gg.
|
||||
have ff:=flimitpres f (imageSet g S). rewrite -> ff. clear gg ff.
|
||||
apply: inf_ext. move => x. unfold imageSet. split.
|
||||
- case => t1. case. case => t0. case => iS gt0 ft1. exists t0. split ; first by []. rewrite <- ft1.
|
||||
by rewrite <- gt0.
|
||||
- case => t0. case => iS fg0. exists (g t0). split ; last by []. exists t0. by split.
|
||||
Qed.
|
||||
|
||||
Canonical Structure mk_flimitM (D0 D1:clatType) (f:fmono D0 D1) (c:limitpres (FMon.Pack (FMon.class f) f)) := flimitMixin c.
|
||||
Definition mk_flimit (D0 D1:clatType) (f:fmono D0 D1) (c:limitpres (FMon.Pack (FMon.class f) f)) := Eval hnf in @flimitType D0 D1 f (mk_flimitM c).
|
||||
|
||||
Definition complp (D0 D1 D2 : clatType) (f : flimit D1 D2) (g: flimit D0 D1) := Eval hnf in mk_flimit (fcomplp f g).
|
||||
|
||||
Lemma oidlp (T:clatType) : limitpres (@oid T).
|
||||
move => S. simpl. apply: inf_ext.
|
||||
move => x. split.
|
||||
- move => iS. exists x. by split. unfold imageSet.
|
||||
- case => y [iS P]. simpl in P. rewrite <- P. by apply iS.
|
||||
Qed.
|
||||
|
||||
Definition lpid (T:clatType) := Eval hnf in mk_flimit (@oidlp T).
|
||||
|
||||
Lemma clatCatAxiom : @Category.axiom _ clatMorphSetoidType complp lpid.
|
||||
split ; last split ; last split.
|
||||
move => T T' f. by apply: fmon_eq_intro.
|
||||
move => T T' f. by apply: fmon_eq_intro.
|
||||
move => T0 T1 T2 T3 f g h. by apply: fmon_eq_intro.
|
||||
move => T0 T1 T2 f f' g g' e e'.
|
||||
apply: fmon_eq_intro => a. simpl. rewrite -> e. by rewrite -> e'.
|
||||
Qed.
|
||||
|
||||
Canonical Structure clatCatMixin := CatMixin clatCatAxiom.
|
||||
Canonical Structure clatCat := Eval hnf in CatType clatCatMixin.
|
||||
|
||||
Section KnasterTarski.
|
||||
|
||||
Section Lat.
|
||||
Variable L : clatType.
|
||||
|
||||
Definition aboveSet : (L -> Prop) -> L -> Prop := (fun D e => forall s, D s -> s <= e).
|
||||
|
||||
Definition sup (S: L -> Prop) : L := inf (aboveSet S).
|
||||
|
||||
Lemma Psup : forall (S : L -> Prop) t, S t -> t <= sup S.
|
||||
intros S t St. unfold sup.
|
||||
refine (PinfL _).
|
||||
intros l Cl. apply Cl. by apply St.
|
||||
Qed.
|
||||
|
||||
Lemma PsupL : forall (S: L -> Prop) t, (forall x, S x -> x <= t) -> sup S <= t.
|
||||
intros S t C. unfold sup. refine (Pinf _).
|
||||
intros l Sl. apply C. apply Sl.
|
||||
Qed.
|
||||
|
||||
|
||||
(** A complete lattice has both infima and suprema. Suprema of a set S ([sup S]) is given by [inf { z | S <= z}] *)
|
||||
|
||||
End Lat.
|
||||
|
||||
Section SubLattice.
|
||||
Variable L : clatType.
|
||||
|
||||
(** Let [P] represent a subset of [L]. *)
|
||||
Variable P : L -> Prop.
|
||||
|
||||
|
||||
|
||||
(** We can then define a an ordering on the subset [P] by inheriting the ordering from [L]. *)
|
||||
|
||||
(** If [P] is closed under infima then it is a complete lattice *)
|
||||
Variable PSinf: forall (S:L -> Prop), (forall t, S t -> P t) -> P (inf S).
|
||||
|
||||
Definition Embed (t:L) (p:P t) : sub_ordType P.
|
||||
by exists t.
|
||||
Defined.
|
||||
|
||||
Definition Extend (S : sub_ordType P -> Prop) : L -> Prop := (fun l => (exists p:P l, S (Embed p))).
|
||||
|
||||
Lemma ExtendP: forall S t, Extend S t -> P t.
|
||||
intros S t Ex. unfold Extend in Ex. unfold Embed in Ex. destruct Ex as [E _].
|
||||
apply E.
|
||||
Defined.
|
||||
|
||||
Lemma sublatAxiom : CLattice.axiom (fun (S:sub_ordType P -> Prop) =>
|
||||
@Embed (inf (Extend S)) (@PSinf (Extend S) (@ExtendP S))).
|
||||
intros S t. split.
|
||||
- case t. simpl. clear t. intros l Pl Pe. refine (Pinf _ ).
|
||||
unfold Extend. unfold Embed. exists Pl. by apply Pe.
|
||||
- case t. simpl. clear t. intros x Px C. apply: PinfL.
|
||||
intros t et. unfold Extend in et. unfold Embed in et.
|
||||
destruct et as [Pt St]. specialize (C _ St). by apply C.
|
||||
Qed.
|
||||
|
||||
Canonical Structure sub_clatMixin := ClatMixin sublatAxiom.
|
||||
Canonical Structure sub_clatType := Eval hnf in ClatType sub_clatMixin.
|
||||
|
||||
End SubLattice.
|
||||
|
||||
Variable L : clatType.
|
||||
|
||||
Lemma op_ordAxiom (O:ordType) : PreOrd.axiom (fun (x y:O) => y <= x).
|
||||
move => x.
|
||||
split ; first by apply Ole_refl.
|
||||
move => y z ; by apply (fun X Y => Ole_trans Y X).
|
||||
Qed.
|
||||
|
||||
Definition op_ordMixin O := OrdMixin (@op_ordAxiom O).
|
||||
Definition op_ordType O := Eval hnf in OrdType (op_ordMixin O).
|
||||
|
||||
Lemma oplatAxiom : @CLattice.axiom (op_ordType L) (@sup L).
|
||||
intros P t. split.
|
||||
- move => Pt. simpl. by apply (Psup Pt).
|
||||
- intros C. simpl. refine (PsupL _). by apply C.
|
||||
Qed.
|
||||
|
||||
Definition op_latMixin := ClatMixin oplatAxiom.
|
||||
Definition op_latType := Eval hnf in @ClatType (op_ordType L) op_latMixin.
|
||||
|
||||
Variable f : ordCatType L L.
|
||||
|
||||
Definition PreFixedPoint d := f d <= d.
|
||||
|
||||
Lemma PreFixedInfs: (forall S : L -> Prop,
|
||||
(forall t : L, S t -> PreFixedPoint t) -> PreFixedPoint (inf S)).
|
||||
intros S C.
|
||||
unfold PreFixedPoint in *.
|
||||
apply (PinfL ). intros l Sl.
|
||||
apply Ole_trans with (y:= f l).
|
||||
assert (monotonic f) as M by auto. apply M.
|
||||
apply Pinf. apply Sl.
|
||||
apply (C _ Sl).
|
||||
Qed.
|
||||
|
||||
Definition PreFixedLattice := @sub_clatType L (fun x => PreFixedPoint x) PreFixedInfs.
|
||||
|
||||
Definition PostFixedPoint d := d <= f d.
|
||||
|
||||
Lemma PostFixedSups: forall S : op_latType -> Prop,
|
||||
(forall t : op_latType, S t -> (fun x : op_latType => PostFixedPoint x) t) ->
|
||||
(fun x : op_latType => PostFixedPoint x) (@inf op_latType S).
|
||||
unfold PostFixedPoint.
|
||||
intros S C. simpl in *. refine (PsupL _).
|
||||
intros l Sl. apply Ole_trans with (y:= f l). apply C. apply Sl.
|
||||
assert (monotonic f) as M by auto. apply M. apply (Psup Sl).
|
||||
Qed.
|
||||
|
||||
Definition PostFixedLattice := @sub_clatType op_latType (fun x => PostFixedPoint x) PostFixedSups.
|
||||
|
||||
Definition lfp := match (@inf PreFixedLattice (fun _ => True)) with exist x _ => x end.
|
||||
|
||||
Definition gfp := match (@inf PostFixedLattice (fun _ => True)) with exist x _ => x end.
|
||||
|
||||
Lemma lfp_fixedpoint : f lfp =-= lfp.
|
||||
unfold lfp.
|
||||
assert (yy:= @Pinf PreFixedLattice (fun _ => True)).
|
||||
generalize yy ; clear yy.
|
||||
case (@inf PreFixedLattice (fun _ => True)).
|
||||
intros x Px yy.
|
||||
assert (zz:=fun X PX => (yy (exist (fun t => PreFixedPoint t) X PX))). clear yy.
|
||||
assert (PreFixedPoint (f x)). unfold PreFixedPoint.
|
||||
assert (monotonic f) as M by auto. apply M. apply Px.
|
||||
split ; first by apply Px.
|
||||
simpl in zz. specialize (zz _ H I). by apply zz.
|
||||
Qed.
|
||||
|
||||
Lemma gfp_fixedpoint : f gfp =-= gfp.
|
||||
unfold gfp.
|
||||
assert (yy:= @Pinf PostFixedLattice (fun _ => True)).
|
||||
generalize yy ; clear yy.
|
||||
case (@inf PostFixedLattice (fun _ => True)).
|
||||
intros x Px yy.
|
||||
assert (zz:=fun X PX => (yy (exist (fun t => PostFixedPoint t) X PX))). clear yy.
|
||||
assert (PostFixedPoint (f x)). unfold PostFixedPoint.
|
||||
assert (monotonic f) as M by auto. apply M. apply Px.
|
||||
split.
|
||||
specialize (zz _ H). simpl in zz. apply zz. auto.
|
||||
apply Px.
|
||||
Qed.
|
||||
|
||||
Lemma lfp_least: forall fp, f fp <= fp -> lfp <= fp.
|
||||
unfold lfp.
|
||||
intros fp ff. assert (PreFixedPoint fp). apply ff.
|
||||
assert (yy:= @Pinf PreFixedLattice (fun _ => True)).
|
||||
generalize yy ; clear yy.
|
||||
case (@inf PreFixedLattice (fun _ => True)).
|
||||
intros x Px yy.
|
||||
assert (zz:=fun X PX => (yy (exist (fun t => PreFixedPoint t) X PX))). clear yy.
|
||||
case (@inf PreFixedLattice (fun _ : PreFixedLattice => True)).
|
||||
specialize (zz _ H). simpl in zz. intros _ _. by apply zz.
|
||||
Qed.
|
||||
|
||||
Lemma gfp_greatest: forall fp, fp <= f fp -> fp <= gfp.
|
||||
unfold gfp.
|
||||
intros fp ff. assert (PostFixedPoint fp). apply ff.
|
||||
assert (yy:= @Pinf PostFixedLattice (fun _ => True)).
|
||||
generalize yy ; clear yy.
|
||||
case (@inf PostFixedLattice (fun _ => True)).
|
||||
intros x Px yy.
|
||||
assert (zz:=fun X PX => (yy (exist (fun t => PostFixedPoint t) X PX))). clear yy.
|
||||
case (@inf PostFixedLattice (fun _ : PostFixedLattice => True)).
|
||||
specialize (zz _ H). simpl in zz. intros _ _. by apply zz.
|
||||
Qed.
|
||||
|
||||
End KnasterTarski.
|
||||
|
||||
Section ClatProd.
|
||||
|
||||
Variable L1 L2 : clatType.
|
||||
|
||||
Lemma prodCLatAxiom : CLattice.axiom (fun (S:L1 * L2 -> Prop) =>
|
||||
(@inf L1 (fun l1 => exists l2, S (l1,l2)),@inf L2 (fun l2 => exists l1, S (l1,l2)))).
|
||||
move => P t. split.
|
||||
- case t. clear t. intros t0 t1 Pt. simpl. split.
|
||||
* refine (Pinf _). exists t1. by apply Pt.
|
||||
* refine (Pinf _). exists t0. by apply Pt.
|
||||
- case t. clear t. intros t0 t1 Pt. simpl. split.
|
||||
* refine (PinfL _). intros x Px. destruct Px as [y Pxy]. specialize (Pt _ Pxy).
|
||||
inversion Pt. simpl in *. by auto.
|
||||
* refine (PinfL _). intros x Px. destruct Px as [y Pyx].
|
||||
specialize (Pt _ Pyx). inversion Pt. simpl in *. by auto.
|
||||
Qed.
|
||||
|
||||
Canonical Structure prod_clatMixin := ClatMixin prodCLatAxiom.
|
||||
Canonical Structure prod_clatType := Eval hnf in ClatType prod_clatMixin.
|
||||
|
||||
Lemma Fst_lp : limitpres (@Fst L1 L2).
|
||||
move => S. unfold imageSet. simpl. apply: inf_ext.
|
||||
move => x. split ; case.
|
||||
- move => y iS. exists (x,y). by split.
|
||||
- case => x' y. simpl. case => iS e. exists y. by rewrite <- e.
|
||||
Qed.
|
||||
|
||||
Definition lfst := Eval hnf in mk_flimit Fst_lp.
|
||||
|
||||
Lemma Snd_lp : limitpres (@Snd L1 L2).
|
||||
move => S. unfold imageSet. simpl. apply: inf_ext.
|
||||
move => x. split ; case.
|
||||
- move => y iS. exists (y,x). by split.
|
||||
- case => x' y. simpl. case => iS e. exists x'. by rewrite <- e.
|
||||
Qed.
|
||||
|
||||
Definition lsnd := Eval hnf in mk_flimit Snd_lp.
|
||||
|
||||
Variable L:clatType.
|
||||
Variable (f:L =-> L1) (g:L =-> L2).
|
||||
|
||||
Lemma prod_fun_lp : limitpres (<|f:ordCatType _ _,g|>).
|
||||
move => S. simpl. rewrite -> (pair_eq_compat (flimitpres f S) (flimitpres g S)).
|
||||
case_eq (inf (imageSet (<| (f : ordCatType _ _), g |>) S)).
|
||||
move => l0 l1 e. unfold imageSet in e. simpl in e. apply: pair_eq_compat.
|
||||
- have aa:=f_equal (Fst _ _) e. simpl in aa. rewrite <- aa. apply: inf_ext. move => x. unfold imageSet.
|
||||
split ; case.
|
||||
+ move => x1 [iS ee]. exists (g x1). exists x1. split ; first by []. by rewrite ee.
|
||||
+ move => x3. case => x1 [iS ee]. exists x1. by split ; case: ee.
|
||||
- have aa:=f_equal (Snd _ _) e. simpl in aa. rewrite <- aa. apply: inf_ext. move => x. unfold imageSet.
|
||||
split ; case.
|
||||
+ move => x1 [iS ee]. exists (f x1). exists x1. split ; first by []. by rewrite ee.
|
||||
+ move => x3. case => x1 [iS ee]. exists x1. by split ; case: ee.
|
||||
Qed.
|
||||
|
||||
Definition lprod_fun : L =-> prod_clatType :=
|
||||
Eval hnf in mk_flimit prod_fun_lp.
|
||||
|
||||
End ClatProd.
|
||||
|
||||
Lemma clatProdCatAxiom : CatProduct.axiom lfst lsnd lprod_fun.
|
||||
move => X Y Z f g. split ; first split.
|
||||
- by apply: fmon_eq_intro.
|
||||
- by apply: fmon_eq_intro.
|
||||
- move => m A. apply: fmon_eq_intro => x.
|
||||
simpl. have A0:=fmon_eq_elim (proj1 A) x. have A1:=fmon_eq_elim (proj2 A) x. by rewrite <- (pair_eq_compat A0 A1).
|
||||
Qed.
|
||||
|
||||
Canonical Structure prod_clatCatMixin := prodCatMixin clatProdCatAxiom.
|
||||
Canonical Structure prod_clatCatType := Eval hnf in prodCatType prod_clatCatMixin.
|
||||
|
||||
Lemma arrow_latAxiom (T:Type) (L:clatType) : @CLattice.axiom (ford_ordType T L)
|
||||
(fun (S:ford_ordType T L -> Prop) => fun t => @inf L (fun st => exists s, S s /\ s t = st)).
|
||||
intros P f. split.
|
||||
- move => Pf. simpl. intros t. refine (Pinf _ ). exists f. by split.
|
||||
- simpl. intros C t. refine (PinfL _).
|
||||
intros l Tl. destruct Tl as [f0 [Pf0 f0t]]. specialize (C _ Pf0). by rewrite <- f0t.
|
||||
Qed.
|
||||
|
||||
Canonical Structure arrow_clatMixin (T:Type) (L:clatType) := ClatMixin (@arrow_latAxiom T L).
|
||||
Canonical Structure arrow_clatType (T:Type) (L:clatType) := Eval hnf in @ClatType (ford_ordType T L) (@arrow_clatMixin T L).
|
||||
|
||||
Require Export NSetoid.
|
||||
Open Scope C_scope.
|
||||
|
||||
Lemma setOrdAxiom T : @PreOrd.axiom (T =-> Props) (fun A B => forall x, A x -> B x).
|
||||
move => c. split ; first by [].
|
||||
move => y z A B a Ca. apply B. by apply A.
|
||||
Qed.
|
||||
|
||||
Canonical Structure set_ordMixin T := OrdMixin (@setOrdAxiom T).
|
||||
Definition set_ordType T := Eval hnf in OrdType (set_ordMixin T).
|
||||
|
||||
Lemma intersect_respect T (S:(T =-> Props) -> Prop) : setoid_respect (fun t, forall A, S A -> A t).
|
||||
move => x y e. simpl.
|
||||
split => X A sa ; specialize (X A sa).
|
||||
- by apply (proj1 (frespect A e)).
|
||||
- by apply (proj2 (frespect A e)).
|
||||
Qed.
|
||||
|
||||
Definition intersect T (S:(T =-> Props) -> Prop) : T =-> Props := Eval hnf in mk_fset (intersect_respect S).
|
||||
|
||||
Lemma set_clatAxiom T : @CLattice.axiom (set_ordType T) (fun S => intersect S).
|
||||
move => P x. split.
|
||||
- move => iP a A. by apply (A x).
|
||||
- move => A a ix B Pb. by apply (A B Pb a ix).
|
||||
Qed.
|
||||
|
||||
Canonical Structure set_clatMixin T := ClatMixin (@set_clatAxiom T).
|
||||
Definition set_clatType T := Eval hnf in @ClatType (@set_ordType T) (@set_clatMixin T).
|
||||
|
||||
Definition set_ordSupType T := Eval hnf in op_ordType (set_ordType T).
|
||||
Definition set_clatSubType T := Eval hnf in op_latType (@set_clatType T).
|
||||
|
||||
|
||||
|
||||
|
1049
papers/domains-2012/MetricCore.v
Normal file
1049
papers/domains-2012/MetricCore.v
Normal file
File diff suppressed because it is too large
Load diff
798
papers/domains-2012/MetricRec.v
Executable file
798
papers/domains-2012/MetricRec.v
Executable file
|
@ -0,0 +1,798 @@
|
|||
(**********************************************************************************
|
||||
* MetricRec.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Categories enriched over bounded complete ultrametric spaces and solving domain
|
||||
equations therein
|
||||
*)
|
||||
|
||||
Require Export Categories MetricCore NSetoid.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
(** printing T0 %\ensuremath{T_0}% *)
|
||||
(** printing T1 %\ensuremath{T_1}% *)
|
||||
(** printing T2 %\ensuremath{T_2}% *)
|
||||
(** printing T3 %\ensuremath{T_3}% *)
|
||||
(** printing T4 %\ensuremath{T_4}% *)
|
||||
(** printing T5 %\ensuremath{T_5}% *)
|
||||
(** printing T6 %\ensuremath{T_6}% *)
|
||||
|
||||
Open Scope C_scope.
|
||||
|
||||
(*=BaseCat *)
|
||||
Module BaseCat.
|
||||
|
||||
Record mixin_of (C:catType) := Mixin
|
||||
{ terminal :> CatTerminal.mixin_of C;
|
||||
metric : forall (X Y : C),
|
||||
Metric.mixin_of (Setoid.Pack (Setoid.class (C X Y)) (C X Y));
|
||||
cmetric : forall (X Y : C),
|
||||
CMetric.mixin_of (Metric.Pack (Metric.Class (metric X Y)) (C X Y)) ;
|
||||
comp: forall (X Y Z : C), (CMetric.Pack (CMetric.Class (cmetric Y Z)) (C Y Z)) *
|
||||
(CMetric.Pack (CMetric.Class (cmetric X Y)) (C X Y)) =->
|
||||
(CMetric.Pack (CMetric.Class (cmetric X Z)) (C X Z));
|
||||
comp_comp : forall X Y Z m m', (comp X Y Z (m,m'):C _ _) =-= m << m' }.
|
||||
(*=End *)
|
||||
|
||||
Section ClassDef.
|
||||
Record class_of (T:Type) (M:T -> T -> setoidType) : Type :=
|
||||
Class { base : Category.class_of M;
|
||||
ext: mixin_of (Category.Pack base T) }.
|
||||
Local Coercion base : class_of >-> Category.class_of.
|
||||
Local Coercion ext : class_of >-> mixin_of.
|
||||
|
||||
Definition base2 (T:Type) (M:T -> T -> setoidType) (c:class_of M) := CatTerminal.Class c.
|
||||
Local Coercion base2 : class_of >-> CatTerminal.class_of.
|
||||
|
||||
Structure cat : Type := Pack {object : Type; morph : object -> object -> setoidType ; _ : class_of morph; _ : Type}.
|
||||
Local Coercion object : cat >-> Sortclass.
|
||||
Local Coercion morph : cat >-> Funclass.
|
||||
Definition class cT := let: Pack _ _ c _ := cT return class_of (@morph cT) in c.
|
||||
Definition unpack (K:forall O (M:O -> O -> setoidType) (c:class_of M), Type)
|
||||
(k : forall O (M: O -> O -> setoidType) (c : class_of M), K _ _ c) (cT:cat) :=
|
||||
let: Pack _ M c _ := cT return @K _ _ (class cT) in k _ _ c.
|
||||
Definition repack cT : _ -> Type -> cat := let k T M c p := p c in unpack k cT.
|
||||
Definition pack := let k T M c m := Pack (@Class T M c m) T in Category.unpack k.
|
||||
|
||||
Definition terminalCat (cT:cat) : terminalCat := CatTerminal.Pack (class cT) cT.
|
||||
Definition catType (cT:cat) : catType := Category.Pack (class cT) cT.
|
||||
Definition metricType (cT:cat) (X Y:cT) : metricType := Metric.Pack (Metric.Class (metric (class cT) X Y)) (cT X Y).
|
||||
Definition cmetricType (cT:cat) (X Y:cT) : cmetricType := CMetric.Pack (CMetric.Class (cmetric (class cT) X Y)) (cT X Y).
|
||||
End ClassDef.
|
||||
Module Import Exports.
|
||||
Coercion terminalCat : cat >-> CatTerminal.cat.
|
||||
Coercion base : class_of >-> Category.class_of.
|
||||
Coercion ext : class_of >-> mixin_of.
|
||||
Coercion object : cat >-> Sortclass.
|
||||
Coercion morph : cat >-> Funclass.
|
||||
Canonical Structure terminalCat.
|
||||
Canonical Structure catType.
|
||||
Canonical Structure metricType.
|
||||
Canonical Structure cmetricType.
|
||||
|
||||
Notation cmetricMorph := cmetricType.
|
||||
Notation cmetricECat := cat.
|
||||
Notation CmetricECatMixin := Mixin.
|
||||
Notation CmetricECatType := pack.
|
||||
End Exports.
|
||||
End BaseCat.
|
||||
Export BaseCat.Exports.
|
||||
|
||||
Definition ccomp (cT:cmetricECat) (X Y Z : cT) : cmetricMorph Y Z * cmetricMorph X Y =-> cmetricMorph X Z :=
|
||||
BaseCat.comp (BaseCat.class cT) X Y Z.
|
||||
|
||||
Lemma ccomp_eq (cT:cmetricECat) (X Y Z : cT) (f : cT Y Z) (f' : cT X Y) : ccomp X Y Z (f,f') =-= f << f'.
|
||||
apply (@BaseCat.comp_comp cT (BaseCat.class cT) X Y Z f f').
|
||||
Qed.
|
||||
|
||||
Section Msolution.
|
||||
Variable M:cmetricECat.
|
||||
|
||||
Open Scope C_scope.
|
||||
Open Scope S_scope.
|
||||
|
||||
(*=BiFunctor *)
|
||||
Record BiFunctor : Type := mk_functor
|
||||
{ ob : M -> M -> M ;
|
||||
morph : forall (T0 T1 T2 T3 : M),
|
||||
(cmetricMorph T1 T0) * (cmetricMorph T2 T3) =->
|
||||
(cmetricMorph (ob T0 T2) (ob T1 T3)) ;
|
||||
morph_comp: forall T0 T1 T2 T3 T4 T5
|
||||
(f:cmetricMorph T4 T1) (g:cmetricMorph T3 T5)
|
||||
(h:cmetricMorph T1 T0) (k:cmetricMorph T2 T3),
|
||||
(morph T1 T4 T3 T5 (f,g)) << morph T0 T1 T2 T3 (h, k)
|
||||
=-= morph _ _ _ _ (h << f, g << k) ;
|
||||
morph_id : forall T0 T1, morph T0 _ T1 _ (Id , Id) =-= Id}.
|
||||
|
||||
(*=End *)
|
||||
|
||||
(*=Tower *)
|
||||
Definition retract (T0 T1 : M) (f: T0 =-> T1) (g: T1 =-> T0) := g << f =-= Id.
|
||||
|
||||
Record Tower : Type := mk_tower
|
||||
{ tobjects : nat -> M;
|
||||
tmorphisms : forall i, tobjects (S i) =-> (tobjects i);
|
||||
tmorphismsI : forall i, (tobjects i) =-> (tobjects (S i));
|
||||
tretract : forall i, retract (tmorphismsI i) (tmorphisms i);
|
||||
tlimitD : forall n i, n <= i ->
|
||||
(tmorphismsI i << tmorphisms i : cmetricMorph _ _) = n = Id}.
|
||||
|
||||
(*=End *)
|
||||
|
||||
(*=Cone *)
|
||||
Record Cone (To:Tower) : Type := mk_basecone
|
||||
{ tcone :> M;
|
||||
mcone : forall i, tcone =-> (tobjects To i);
|
||||
mconeCom : forall i, tmorphisms To i << mcone (S i) =-= mcone i }.
|
||||
(*=End *)
|
||||
|
||||
Record CoCone (To:Tower) : Type := mk_basecocone
|
||||
{ tcocone :> M;
|
||||
mcocone : forall i, (tobjects To i) =-> tcocone;
|
||||
mcoconeCom : forall i, mcocone (S i) << tmorphismsI To i =-= mcocone i }.
|
||||
|
||||
(*=Limit *)
|
||||
Record Limit (To:Tower) : Type := mk_baselimit
|
||||
{ lcone :> Cone To;
|
||||
limitExists : forall (A:Cone To), (tcone A) =-> (tcone lcone);
|
||||
limitCom : forall (A:Cone To), forall n, mcone A n =-= mcone lcone n << limitExists A;
|
||||
limitUnique : forall (A:Cone To) (h: (tcone A) =-> (tcone lcone))
|
||||
(C:forall n, mcone A n =-= mcone lcone n << h), h =-= limitExists A }.
|
||||
(*=End *)
|
||||
Record CoLimit (To:Tower) : Type := mk_basecolimit
|
||||
{ lcocone :> CoCone To;
|
||||
colimitExists : forall (A:CoCone To), (tcocone lcocone) =-> (tcocone A) ;
|
||||
colimitCom : forall (A:CoCone To), forall n, mcocone A n =-= colimitExists A << mcocone lcocone n;
|
||||
colimitUnique : forall (A:CoCone To) (h: (tcocone lcocone) =-> (tcocone A))
|
||||
(C:forall n, mcocone A n =-= h << mcocone lcocone n), h =-= colimitExists A }.
|
||||
|
||||
Variable L:forall T:Tower, Limit T.
|
||||
|
||||
Section RecursiveDomains.
|
||||
|
||||
Variable F : BiFunctor.
|
||||
|
||||
Add Parametric Morphism C X Y Z W : (fun (x: _ =-> _) (y: _ =-> _) => @morph C X Y Z W (pair x y))
|
||||
with signature (@tset_eq _) ==> (@tset_eq _) ==> (@tset_eq _)
|
||||
as morph_eq_compat.
|
||||
move => x x' e y y' e'. apply: (frespect (morph C X Y Z W)). by split.
|
||||
Qed.
|
||||
|
||||
Lemma BiFuncRtoR: forall (T0 T1:M) (f: T0 =-> T1) (g: T1 =-> T0), retract f g ->
|
||||
retract (morph F _ _ _ _ (g,f)) (morph F _ _ _ _ (f,g)).
|
||||
move => A B f g X. unfold retract.
|
||||
- rewrite -> morph_comp. rewrite -> (morph_eq_compat F X X).
|
||||
by rewrite morph_id.
|
||||
Qed.
|
||||
|
||||
Section Tower.
|
||||
|
||||
Variable DTower : Tower.
|
||||
|
||||
Lemma lt_gt_dec n m : (n < m) + (m = n) + (m < n).
|
||||
case: (ltngtP m n) => e.
|
||||
by right.
|
||||
left. by left.
|
||||
left. by right.
|
||||
Qed.
|
||||
|
||||
Definition Diter_coercion n m (Eq:n = m) : M (tobjects DTower n) (tobjects DTower m).
|
||||
rewrite Eq. by apply: Id.
|
||||
Defined.
|
||||
|
||||
Lemma leqI m n : (m <= n)%N = false -> n <= m.
|
||||
move => I. have t:=(leq_total m n). rewrite I in t. by apply t.
|
||||
Qed.
|
||||
|
||||
Lemma lt_subK n m : m < n -> n = (n - m + m)%N.
|
||||
move => l. apply sym_eq. apply subnK. have x:= (leqnSn m). by apply (leq_trans x l).
|
||||
Qed.
|
||||
|
||||
Fixpoint Projection_nm m (n:nat) : (tobjects DTower (n+m)) =-> (tobjects DTower m) :=
|
||||
match n as n0 return (tobjects DTower (n0+m)) =-> (tobjects DTower m) with
|
||||
| O => Id
|
||||
| S n => Projection_nm m n << tmorphisms DTower (n+m)%N
|
||||
end.
|
||||
|
||||
Fixpoint Injection_nm m (n:nat) : (tobjects DTower m) =-> (tobjects DTower (n+m)) :=
|
||||
match n as n0 return (tobjects DTower m) =-> (tobjects DTower (n0+m)) with
|
||||
| O => Id
|
||||
| S n => tmorphismsI DTower (n+m) << Injection_nm m n
|
||||
end.
|
||||
|
||||
Definition t_nm n m : (tobjects DTower n) =-> (tobjects DTower m) :=
|
||||
match (lt_gt_dec m n) return (tobjects DTower n) =-> (tobjects DTower m) with
|
||||
| inl (inl ee) => (Projection_nm m (n - m)%N) << (Diter_coercion (lt_subK ee))
|
||||
| inl (inr ee) => Diter_coercion ee : (tobjects DTower n) =-> (tobjects DTower m)
|
||||
| inr ee => Diter_coercion (sym_eq (lt_subK ee)) << Injection_nm n (m - n)
|
||||
end.
|
||||
|
||||
Lemma Diter_coercion_simpl: forall n (Eq:n = n), Diter_coercion Eq = Id.
|
||||
intros n Eq. by rewrite (eq_irrelevance Eq (refl_equal _)).
|
||||
Qed.
|
||||
|
||||
Lemma Proj_left_comp : forall k m, Projection_nm m (S k) =-=
|
||||
tmorphisms DTower m << Projection_nm (S m) k << Diter_coercion (sym_eq (addnS k m)).
|
||||
elim.
|
||||
- move => m. simpl.
|
||||
rewrite comp_idL. rewrite comp_idR.
|
||||
rewrite (Diter_coercion_simpl). by rewrite -> comp_idR.
|
||||
- move => n IH m. simpl. simpl in IH. rewrite -> IH.
|
||||
have P:=(sym_eq (addnS n m)). rewrite (eq_irrelevance (sym_eq (addnS n m)) P).
|
||||
have Q:=(sym_eq (addnS n.+1 m)). rewrite (eq_irrelevance (sym_eq (addnS n.+1 m)) Q).
|
||||
move: P Q. unfold addn. simpl. fold addn. move => P. generalize P. rewrite P. clear P.
|
||||
move => P Q. do 2 rewrite Diter_coercion_simpl.
|
||||
do 2 rewrite comp_idR. clear P Q. by rewrite comp_assoc.
|
||||
Qed.
|
||||
|
||||
Lemma Inject_right_comp: forall k m, Injection_nm m (S k) =-=
|
||||
Diter_coercion (addnS k m) << Injection_nm (S m) k << tmorphismsI DTower m.
|
||||
elim.
|
||||
- move => m. simpl. do 2 rewrite comp_idR. rewrite (Diter_coercion_simpl). by rewrite -> comp_idL.
|
||||
- move => n IH m. simpl. simpl in IH. rewrite -> IH. clear IH.
|
||||
have P:= (addnS n.+1 m). have Q:=(addnS n m). rewrite -> (eq_irrelevance _ P).
|
||||
rewrite -> (eq_irrelevance _ Q). move: P Q. unfold addn. simpl. fold addn.
|
||||
move => P Q. generalize P Q. rewrite <- Q. clear P Q. move => P Q. do 2 rewrite Diter_coercion_simpl.
|
||||
do 2 rewrite comp_idL. by rewrite comp_assoc.
|
||||
Qed.
|
||||
|
||||
Lemma Diter_coercion_comp: forall x y z (Eq1:x = y) (Eq2 : y = z),
|
||||
Diter_coercion Eq2 << Diter_coercion Eq1 =-= Diter_coercion (trans_equal Eq1 Eq2).
|
||||
intros x y z Eq1. generalize Eq1. rewrite Eq1. clear x Eq1.
|
||||
intros Eq1 Eq2. generalize Eq2. rewrite <- Eq2. clear Eq2. intros Eq2.
|
||||
repeat (rewrite Diter_coercion_simpl). by rewrite -> comp_idL.
|
||||
Qed.
|
||||
|
||||
Lemma Diter_coercion_eq n m (e e': n = m) : Diter_coercion e =-= Diter_coercion e'.
|
||||
by rewrite (eq_irrelevance e e').
|
||||
Qed.
|
||||
|
||||
Lemma lt_gt_dec_lt n m : n < m -> exists e, lt_gt_dec n m = inl _ (inl _ e).
|
||||
move => e. case_eq (lt_gt_dec n m) ; first case.
|
||||
- move => i _. by exists i.
|
||||
- move => f. rewrite f in e. by rewrite ltnn in e.
|
||||
- move => f. have a:=leq_trans (leqnSn _) f. have Fx:=leq_ltn_trans a e. by rewrite ltnn in Fx.
|
||||
Qed.
|
||||
|
||||
Lemma lt_gt_dec_gt n m : m < n -> exists e, lt_gt_dec n m = (inr _ e).
|
||||
move => e. case_eq (lt_gt_dec n m) ; first case.
|
||||
- move => i _. have a:=leq_trans (leqnSn _) i. have Fx:=leq_ltn_trans a e. by rewrite ltnn in Fx.
|
||||
- move => f. rewrite f in e. by rewrite ltnn in e.
|
||||
- move => i _. by exists i.
|
||||
Qed.
|
||||
|
||||
Lemma lt_gt_dec_eq n m : forall e : m = n, lt_gt_dec n m = (inl _ (inr _ e)).
|
||||
move => e. generalize e. rewrite e. clear e m. move => e. case_eq (lt_gt_dec n n) ; first case.
|
||||
- move => i _. have f:=i. by rewrite ltnn in f.
|
||||
- move => f _. by rewrite (eq_irrelevance e f).
|
||||
- move => i _. have f:=i. by rewrite ltnn in f.
|
||||
Qed.
|
||||
|
||||
Lemma t_nmProjection: forall n m, t_nm n m =-= tmorphisms DTower m << t_nm n (S m).
|
||||
move => n m. unfold t_nm. case (lt_gt_dec m n) ; first case.
|
||||
- case (lt_gt_dec m.+1 n) ; first case.
|
||||
+ move => l l'. rewrite comp_assoc.
|
||||
have X:=Proj_left_comp (n - m.+1) m. have XX:= addnS (n - m.+1) m.
|
||||
have Y:=comp_eq_compat X (tset_refl (Diter_coercion XX)).
|
||||
repeat (rewrite <- comp_assoc in Y). rewrite -> Diter_coercion_comp in Y.
|
||||
rewrite Diter_coercion_simpl in Y. rewrite -> comp_idR in Y. rewrite <- Y.
|
||||
rewrite <- comp_assoc. rewrite Diter_coercion_comp.
|
||||
have P:=(trans_equal (lt_subK l) XX). have Q:=(lt_subK l').
|
||||
rewrite -> (eq_irrelevance _ P). rewrite -> (eq_irrelevance _ Q).
|
||||
generalize P. have e:= (ltn_subS l'). have ee:(n - m).-1 = (n - m.+1). by rewrite e.
|
||||
rewrite <- ee. clear ee e P. have e:0 < n - m. rewrite <- ltn_add_sub. by rewrite addn0.
|
||||
move => P. have PP:n = ((n - m).-1.+1 + m)%N. rewrite (ltn_predK e). apply Q.
|
||||
rewrite (eq_irrelevance P PP). move: PP. rewrite (ltn_predK e). clear. move => PP.
|
||||
by rewrite (eq_irrelevance Q PP).
|
||||
+ move => e. generalize e. rewrite e. clear n e. move => e i. have x:= subSnn m.
|
||||
have a:=(lt_subK (n:=m.+1) (m:=m) i). rewrite (eq_irrelevance (lt_subK (n:=m.+1) (m:=m) i) a).
|
||||
move: a. rewrite x. clear x. move => a. rewrite (eq_irrelevance e a). simpl. by rewrite comp_idL.
|
||||
+ move => i j. have Fx:=leq_trans i j. by rewrite ltnn in Fx.
|
||||
- move => e. generalize e. rewrite e. clear n e. move => e. destruct (lt_gt_dec_gt (leqnn (S m))) as [p A].
|
||||
rewrite A. clear A.
|
||||
have ee:=subSnn m. have x:=(sym_eq (lt_subK (n:=m.+1) (m:=m) p)).
|
||||
rewrite <- (eq_irrelevance x (sym_eq (lt_subK (n:=m.+1) (m:=m) p))). move: x. rewrite ee. simpl.
|
||||
move => x. do 2 rewrite (Diter_coercion_simpl). rewrite comp_idL. rewrite comp_idR.
|
||||
unfold addn. simpl. by rewrite -> (tretract DTower m).
|
||||
- move => i. have l:=leq_trans i (leqnSn m). destruct (lt_gt_dec_gt l) as [p A]. rewrite A. clear A l.
|
||||
have x:=(sym_eq (lt_subK (n:=m.+1) (m:=n) p)). rewrite (eq_irrelevance (sym_eq (lt_subK (n:=m.+1) (m:=n) p)) x).
|
||||
move: x. rewrite leq_subS ; last by []. simpl. move => x. do 2 rewrite comp_assoc.
|
||||
refine (comp_eq_compat _ (tset_refl _)). have y:=(sym_eq (lt_subK (n:=m) (m:=n) i)).
|
||||
rewrite (eq_irrelevance (sym_eq (lt_subK (n:=m) (m:=n) i)) y). have a:=(trans_eq (sym_eq (addSn _ _)) x).
|
||||
rewrite (eq_irrelevance x a).
|
||||
apply tset_trans with (y:=(tmorphisms DTower m << Diter_coercion a) << tmorphismsI DTower (m - n + n)) ; last by [].
|
||||
move: a. generalize y. rewrite -> y. clear. move => y a. do 2 rewrite Diter_coercion_simpl.
|
||||
rewrite comp_idR. by rewrite (tretract DTower m).
|
||||
Qed.
|
||||
|
||||
Lemma t_nn_ID: forall n, t_nm n n =-= Id.
|
||||
intros n.
|
||||
unfold t_nm. rewrite (lt_gt_dec_eq (refl_equal n)). by rewrite Diter_coercion_simpl.
|
||||
Qed.
|
||||
|
||||
Lemma t_nmProjection2 n m : (m <= n) % nat -> t_nm (S n) m =-= (t_nm n m) << tmorphisms DTower n.
|
||||
move => nm.
|
||||
assert (exists x, n - m = x) as X by (exists (n - m) ; auto).
|
||||
destruct X as [x X]. elim: x n m X nm.
|
||||
- move => n m E EE. have e:=eqn_leq n m. rewrite EE in e. have l:(n <= m)%N. unfold leq. by rewrite E.
|
||||
rewrite l in e. have ee:=eqP e. rewrite ee. rewrite t_nn_ID. rewrite comp_idL.
|
||||
rewrite -> t_nmProjection. rewrite t_nn_ID. by rewrite comp_idR.
|
||||
- move => nm IH. case ; first by []. move => n m e l. have ll:n.+1 - m > 0 by rewrite e. rewrite subn_gt0 in ll.
|
||||
rewrite leq_subS in e ; last by []. specialize (IH n m). have ee: n - m = nm by auto.
|
||||
rewrite -> (IH ee ll).
|
||||
clear IH. have l0:=leq_trans ll (leqnSn _). destruct (lt_gt_dec_lt l0) as [p A].
|
||||
unfold t_nm. rewrite A. clear A. case_eq (lt_gt_dec m n) ; first case.
|
||||
+ move => p' _. have x:=(lt_subK (n:=n.+2) (m:=m) p). rewrite (eq_irrelevance (lt_subK (n:=n.+2) (m:=m) p) x).
|
||||
move: x. have x:n.+2 - m = (n-m).+2. rewrite leq_subS ; last by []. rewrite leq_subS ; by [].
|
||||
rewrite x. clear x. move => x. simpl. do 4 rewrite <- comp_assoc. refine (comp_eq_compat (tset_refl _) _).
|
||||
have x':=(lt_subK (n:=n) (m:=m) p'). rewrite (eq_irrelevance (lt_subK (n:=n) (m:=m) p') x').
|
||||
have x1:((n - m).+2 + m)%N = ((n - m) + m).+2 by [].
|
||||
have x2:=trans_eq x x1.
|
||||
apply tset_trans with (y:=tmorphisms DTower (n - m + m) << (tmorphisms DTower (S(n - m + m)) << Diter_coercion x2)) ;
|
||||
first by simpl ; rewrite (eq_irrelevance x2 x).
|
||||
move: x2. generalize x' ; rewrite <- x'. clear x' x x1 p p'. move => x x'. do 2 rewrite (Diter_coercion_simpl).
|
||||
rewrite comp_idR. by rewrite -> comp_idL.
|
||||
+ move => e'. generalize e'. move: p. rewrite e'. clear e' n l ll e l0 ee. move => p e _.
|
||||
rewrite (Diter_coercion_simpl). rewrite comp_idL.
|
||||
have x:=(lt_subK (n:=m.+2) (m:=m) p).
|
||||
have ee:(m.+2 - m) = (m - m).+2. rewrite leq_subS ; last by []. rewrite leq_subS ; by [].
|
||||
rewrite (eq_irrelevance (lt_subK (n:=m.+2) (m:=m) p) x). generalize x. rewrite ee. clear ee. rewrite subnn.
|
||||
move => xx. clear x e p. simpl. rewrite comp_idL.
|
||||
rewrite (Diter_coercion_simpl). by rewrite -> comp_idR.
|
||||
+ move => f. have lx:=leq_trans f ll. by rewrite ltnn in lx.
|
||||
Qed.
|
||||
|
||||
Lemma t_nmEmbedding: forall n i, t_nm n i =-= (t_nm (S n) i) << tmorphismsI DTower n.
|
||||
intros n m. unfold t_nm. case (lt_gt_dec m n) ; first case.
|
||||
- move => i. have l:=leq_trans i (leqnSn _). destruct (lt_gt_dec_lt l) as [p A]. rewrite A. clear A l.
|
||||
have x:=((lt_subK p)). rewrite (eq_irrelevance ((lt_subK p)) x).
|
||||
move: x. rewrite leq_subS ; last by []. simpl. move => x. do 2 rewrite <- comp_assoc.
|
||||
refine (comp_eq_compat (tset_refl _) _). have y:=((lt_subK i)).
|
||||
rewrite (eq_irrelevance ((lt_subK i)) y). have a:=(trans_eq x ((addSn _ _))).
|
||||
rewrite (eq_irrelevance x a).
|
||||
apply tset_trans with (y:=tmorphisms DTower (n - m + m) << (Diter_coercion a << tmorphismsI DTower n)) ; last by [].
|
||||
move: a. generalize y. rewrite <- y. clear. move => y a. do 2 rewrite Diter_coercion_simpl.
|
||||
rewrite comp_idL. by rewrite <- (tretract DTower n).
|
||||
- move => e. generalize e. rewrite e. clear n e. move => e. destruct (lt_gt_dec_lt (leqnn (S m))) as [p A].
|
||||
rewrite A. clear A.
|
||||
have ee:=subSnn m. have x:=((lt_subK (n:=m.+1) (m:=m) p)).
|
||||
rewrite <- (eq_irrelevance x ((lt_subK (n:=m.+1) (m:=m) p))). move: x. rewrite ee. simpl.
|
||||
move => x. do 2 rewrite (Diter_coercion_simpl). rewrite comp_idR. rewrite comp_idL.
|
||||
by rewrite -> (tretract DTower m).
|
||||
- case (lt_gt_dec m n.+1) ; first case.
|
||||
+ move => i j. have Fx:=leq_trans i j. by rewrite ltnn in Fx.
|
||||
+ move => e. generalize e. rewrite <- e. clear m e. move => e i. have x:= subSnn n.
|
||||
have a:=(lt_subK i). rewrite (eq_irrelevance (lt_subK i) a).
|
||||
move: a. rewrite x. clear x. move => a. rewrite (eq_irrelevance (sym_eq a) e).
|
||||
simpl. by rewrite comp_idR.
|
||||
+ move => l l'. rewrite <- comp_assoc. have X:=Inject_right_comp (m - n.+1) n. have XX:= sym_eq (addnS (m - n.+1) n).
|
||||
have Y:=comp_eq_compat (tset_refl (Diter_coercion XX)) X. rewrite -> comp_assoc in Y.
|
||||
rewrite -> (comp_assoc (Injection_nm _ _)) in Y. rewrite -> Diter_coercion_comp in Y.
|
||||
rewrite Diter_coercion_simpl in Y. rewrite -> comp_idL in Y. rewrite <- Y.
|
||||
rewrite comp_assoc. rewrite Diter_coercion_comp.
|
||||
have ee:=(trans_equal XX (sym_eq (lt_subK l))). rewrite -> (eq_irrelevance _ ee).
|
||||
clear XX X Y. have e:m - n.+1 = (m - n).-1. by rewrite (ltn_subS l').
|
||||
move: ee. rewrite -> e. have ll:0 < (m - n). rewrite <- ltn_add_sub. by rewrite addn0.
|
||||
move => Q.
|
||||
have PP:((m - n).-1.+1 + n)%N = m. rewrite -> (ltn_predK ll). by rewrite (sym_eq (lt_subK l')).
|
||||
rewrite (eq_irrelevance Q PP). move: PP. have P:=(sym_eq (lt_subK l')). rewrite -> (eq_irrelevance _ P).
|
||||
move: P. clear Q. clear e. move => P Q. move: P. rewrite -{1 2 3 4} (ltn_predK ll).
|
||||
move => P. by rewrite (eq_irrelevance P Q).
|
||||
Qed.
|
||||
|
||||
Lemma t_nmEmbedding2: forall n m, (n <= m) % nat -> t_nm n (S m) =-= tmorphismsI DTower m << t_nm n m.
|
||||
move => n m nm.
|
||||
assert (exists x, m - n = x) as X by (eexists ; apply refl_equal).
|
||||
destruct X as [x X]. elim: x n m X nm.
|
||||
- move => n m E EE. have e:=eqn_leq n m. rewrite EE in e. have l:(m <= n)%N. unfold leq. by rewrite E.
|
||||
rewrite l in e. have ee:=eqP e. rewrite ee. rewrite -> t_nn_ID. rewrite comp_idR.
|
||||
rewrite -> t_nmEmbedding. rewrite t_nn_ID. by rewrite comp_idL.
|
||||
- move => nm IH. move => n. case ; first by []. move => m e l. have ll:m.+1 - n > 0 by rewrite e. rewrite subn_gt0 in ll.
|
||||
rewrite leq_subS in e ; last by []. specialize (IH n m). have ee:m - n = nm by auto.
|
||||
rewrite (IH ee ll). clear IH. have l0:=leq_trans ll (leqnSn _). destruct (lt_gt_dec_gt l0) as [p A].
|
||||
unfold t_nm. rewrite A. clear A. case_eq (lt_gt_dec m n) ; first case.
|
||||
+ move => f. have lx:=leq_trans f ll. by rewrite ltnn in lx.
|
||||
+ move => e'. generalize e'. move: p. rewrite e'. clear e' n l ll e l0 ee. move => p e _.
|
||||
rewrite (Diter_coercion_simpl). rewrite comp_idR. have x:=(lt_subK p).
|
||||
have ee:(m.+2 - m) = (m - m).+2. rewrite leq_subS ; last by []. rewrite leq_subS ; by [].
|
||||
rewrite (eq_irrelevance (lt_subK (n:=m.+2) (m:=m) p) x). generalize x. rewrite ee. clear ee. rewrite subnn.
|
||||
move => xx. clear x e p. simpl. rewrite (Diter_coercion_simpl). rewrite -> comp_idL. by rewrite comp_idR.
|
||||
+ move => p' _. have x:=(lt_subK p). rewrite (eq_irrelevance (lt_subK p) x).
|
||||
move: x. have x:m.+2 - n = (m-n).+2. rewrite leq_subS ; last by []. rewrite leq_subS ; by [].
|
||||
rewrite x. clear x. move => x. simpl. do 4 rewrite -> comp_assoc. refine (comp_eq_compat _ (tset_refl _)).
|
||||
have x':=sym_eq(lt_subK p'). rewrite (eq_irrelevance (sym_eq (lt_subK _)) x').
|
||||
have x1:((m - n).+2 + n)%N = ((m - n) + n).+2 by [].
|
||||
have x2:=sym_eq (trans_eq x x1).
|
||||
apply tset_trans with (y:=(Diter_coercion x2 << tmorphismsI DTower (S (m - n + n))) << tmorphismsI DTower ((m - n + n))) ;
|
||||
first by simpl ; rewrite (eq_irrelevance x2 (sym_eq x)).
|
||||
move: x2. generalize x' ; rewrite x'. clear x' x x1 p p'. move => x x'. do 2 rewrite (Diter_coercion_simpl).
|
||||
rewrite -> comp_idR. by rewrite comp_idL.
|
||||
Qed.
|
||||
|
||||
Lemma t_nm_EP i n : (i <= n)%N -> retract (t_nm i n) (t_nm n i).
|
||||
have e:exists x, x == n - i by eexists.
|
||||
case: e => x e. elim: x i n e.
|
||||
- move => i n e l. have ee:i = n. apply anti_leq. rewrite l. simpl. unfold leq. by rewrite <- (eqP e).
|
||||
rewrite -> ee. unfold retract. rewrite -> t_nn_ID. by rewrite comp_idL.
|
||||
- move => x IH i. case ; first by rewrite sub0n.
|
||||
move => n e l. have ll:(i <= n)%N. have aa:0 < n.+1 - i. by rewrite <- (eqP e). rewrite subn_gt0 in aa. by apply aa.
|
||||
clear l. specialize (IH i n). rewrite (leq_subS ll) in e. specialize (IH e ll). unfold retract.
|
||||
rewrite (t_nmProjection2 ll). rewrite -> (t_nmEmbedding2 ll). rewrite comp_assoc.
|
||||
rewrite <- (comp_assoc (tmorphismsI _ _)). rewrite (tretract DTower n). rewrite comp_idR. by apply (IH).
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism A B C n : (@Category.comp M A B C)
|
||||
with signature (fun x y => @Mrel (cmetricMorph B C) n x y) ==> (fun x y => @Mrel (cmetricMorph A B) n x y) ==>
|
||||
(fun x y => @Mrel (cmetricMorph A C) n x y)
|
||||
as comp_neq_compat.
|
||||
move => x y e x' y' e'.
|
||||
rewrite <- (ccomp_eq x). rewrite <- (proj2 (Mrefl (ccomp _ _ _ (y,y')) _) (ccomp_eq y y') n).
|
||||
apply: fnonexpansive. by split.
|
||||
Qed.
|
||||
|
||||
Lemma chainPEp (C:CoCone DTower) (C':Cone DTower) : cchainp (fun n : nat => mcocone C n << mcone C' n:cmetricMorph (tcone C') (tcocone C)).
|
||||
move => n i j l l'. have X:= tlimitD DTower.
|
||||
have A: forall k, n <= k -> (mcocone C n << mcone C' n:cmetricMorph (tcone C') (tcocone C)) = n = (mcocone C k << mcone C' k).
|
||||
- elim ; first by move => lx ; rewrite leqn0 in lx ; rewrite (eqP lx).
|
||||
move => x IH lx. case_eq (n <= x) => a.
|
||||
+ specialize (IH a). refine (Mrel_trans IH _).
|
||||
have Cx:=comp_eq_compat (mcoconeCom C x) ((mconeCom C' x)).
|
||||
simpl in Cx. rewrite <- comp_assoc in Cx. rewrite -> (comp_assoc (mcone C' _)) in Cx.
|
||||
refine (Mrel_trans ((proj2 (@Mrefl (cmetricMorph (tcone C') (tcocone C)) _ _) (tset_sym Cx) n)) _). clear Cx. specialize (X _ _ a).
|
||||
apply comp_neq_compat ; first by [].
|
||||
apply Mrel_trans with (y:=Id << mcone C' x.+1) ; last by apply: (proj2 (Mrefl _ _)) ; apply: comp_idL.
|
||||
apply comp_neq_compat ; last by [].
|
||||
by apply X.
|
||||
+ have b:= ltnP x n. rewrite a in b. have c:(x < n) by inversion b. clear b.
|
||||
have b:=@anti_leq n (x.+1). rewrite c in b. rewrite lx in b. specialize (b is_true_true).
|
||||
by rewrite b.
|
||||
have B:=A _ l. specialize (A _ l'). by apply (Mrel_trans (Msym B) A).
|
||||
Qed.
|
||||
|
||||
Definition chainPE (C:CoCone DTower) (C':Cone DTower) : cchain (cmetricMorph (tcone C') (tcocone C)) :=
|
||||
mk_cchain (@chainPEp C C').
|
||||
|
||||
Definition coneN (n:nat) : Cone DTower.
|
||||
exists (tobjects DTower n) (t_nm n).
|
||||
move => m. simpl. by rewrite ->(t_nmProjection n m).
|
||||
Defined.
|
||||
|
||||
Lemma coneCom_l (X:Cone DTower) n i : (i <= n)%nat -> mcone X i =-= t_nm n i << mcone X n.
|
||||
move => l.
|
||||
assert (exists x, n - i = x) as E by (eexists ; auto).
|
||||
destruct E as [x E]. elim: x n i l E.
|
||||
- move => n i l E. have A:=@anti_leq n i. rewrite l in A. unfold leq in A. rewrite E in A.
|
||||
rewrite andbT in A. specialize (A (eqtype.eq_refl _)). rewrite A. clear l E A n.
|
||||
rewrite t_nn_ID. by rewrite comp_idL.
|
||||
- move => x IH n i l E. rewrite -> (comp_eq_compat (t_nmProjection n i) (tset_refl _)).
|
||||
rewrite <- comp_assoc. specialize (IH n i.+1). have l0:i < n by rewrite <- subn_gt0 ; rewrite E.
|
||||
have ee:n - i.+1 = x. have Y:= subn_gt0 i n. rewrite E in Y. rewrite ltn_subS in E ; first by auto. by rewrite <- Y.
|
||||
specialize (IH l0 ee). rewrite <- IH.
|
||||
by rewrite -> (mconeCom X i).
|
||||
Qed.
|
||||
|
||||
Lemma coconeCom_l (X:CoCone DTower) n i : (n <= i)%nat -> mcocone X n =-= mcocone X i << (t_nm n i).
|
||||
move => l.
|
||||
assert (exists x, i - n = x) as E by (eexists ; auto).
|
||||
destruct E as [x E]. elim: x n i l E.
|
||||
- move => n i l E. have A:=@anti_leq n i. rewrite l in A. unfold leq in A. rewrite E in A.
|
||||
specialize (A is_true_true). rewrite A. rewrite -> t_nn_ID. by rewrite comp_idR.
|
||||
- move => x IH n i l E. rewrite -> t_nmEmbedding.
|
||||
rewrite -> comp_assoc. specialize (IH n.+1 i).
|
||||
have ll:n < i by rewrite <- subn_gt0 ; rewrite E.
|
||||
have ee:i - n.+1 = x by have Y:= subn_gt0 n i ; rewrite E in Y ;
|
||||
rewrite ltn_subS in E ; try rewrite <- Y ; by auto.
|
||||
rewrite <- (IH ll ee). by apply (tset_sym (mcoconeCom X n)).
|
||||
Qed.
|
||||
|
||||
End Tower.
|
||||
|
||||
Fixpoint Diter (n:nat) :=
|
||||
match n return M with
|
||||
| O => One
|
||||
| S n => ob F (Diter n) (Diter n)
|
||||
end.
|
||||
|
||||
Variable tmorph_ne : One =-> (ob F One One).
|
||||
|
||||
Fixpoint Injection (n:nat) : (Diter n) =-> (Diter (S n)) :=
|
||||
match n with
|
||||
| O => tmorph_ne
|
||||
| S n => morph F (Diter n) (Diter (S n)) (Diter n) (Diter (S n))
|
||||
(Projection n, Injection n)
|
||||
end
|
||||
with Projection (n:nat) : (Diter (S n)) =-> (Diter n) :=
|
||||
match n with
|
||||
| O => terminal_morph _
|
||||
| S n => morph F (Diter (S n)) (Diter n) (Diter (S n)) (Diter n)
|
||||
(Injection n,Projection n)
|
||||
end.
|
||||
|
||||
Lemma retract_IP: forall n, retract (Injection n) (Projection n).
|
||||
elim ; first by apply: terminal_unique.
|
||||
move => n IH. unfold retract. simpl. rewrite -> morph_comp.
|
||||
unfold retract in IH. rewrite -> (morph_eq_compat F IH IH).
|
||||
by rewrite -> morph_id.
|
||||
Qed.
|
||||
|
||||
Lemma IP_nonexp i j n : i <= j -> (Injection i << Projection i:cmetricMorph _ _) = n = Id -> (Injection j << Projection j : cmetricMorph _ _) = n = Id.
|
||||
move => l.
|
||||
have M':exists m, m = j - i by eexists ; apply refl_equal. destruct M' as [m M'].
|
||||
elim: m i j l M'.
|
||||
- move => i j l e. have E:= (subn_eq0 j i). rewrite e in E. rewrite eqtype.eq_refl in E.
|
||||
have A:= anti_leq. specialize (A i j). rewrite l in A. rewrite <- E in A. specialize (A is_true_true).
|
||||
by rewrite A.
|
||||
- move => m IH i. case ; first by [].
|
||||
move=> j l e. specialize (IH i j). have ll: 0 < j.+1 - i. by rewrite <- e.
|
||||
rewrite subn_gt0 in ll. specialize (IH ll). clear l. move => X. have I:= (IH _ X). clear IH X.
|
||||
simpl. refine (Mrel_trans (proj2 (Mrefl (_ : cmetricMorph _ _) _) (morph_comp F _ _ _ _) n) _).
|
||||
apply: (Mrel_trans _ (proj2 (Mrefl (_ : cmetricMorph _ _) _) (morph_id F _ _) n)).
|
||||
apply: fnonexpansive. split ; apply I ; rewrite leq_subS in e ; by auto.
|
||||
Qed.
|
||||
|
||||
Variable morph_contractive : forall (T0 T1 T2 T3:M), contractive (morph F T0 T1 T2 T3).
|
||||
|
||||
Lemma IP_cauchy n i : n <= i -> (Injection i << Projection i:cmetricMorph (Diter i.+1) (Diter i.+1)) = n = Id.
|
||||
- elim:n i ; first by [].
|
||||
move => n IH. case ; first by []. move => i l. specialize (IH i l).
|
||||
simpl. refine (Mrel_trans (proj2 (Mrefl (_ : cmetricMorph _ _) _) (morph_comp F _ _ _ _) (S n)) _).
|
||||
refine (Mrel_trans _ (proj2 (Mrefl (_ : cmetricMorph _ _) _) (morph_id F _ _) (S n))).
|
||||
apply morph_contractive. simpl. by split.
|
||||
Qed.
|
||||
|
||||
Definition DTower := mk_tower retract_IP IP_cauchy.
|
||||
|
||||
(** printing DInf %\ensuremath{D_\infty}% *)
|
||||
(*=DInf *)
|
||||
Definition DInf : M := tcone (L DTower).
|
||||
(*=End *)
|
||||
Definition Embeddings n : (tobjects DTower n) =-> DInf := limitExists (L DTower) (coneN DTower n).
|
||||
|
||||
Definition Projections n : DInf =-> Diter n := mcone (L DTower) n.
|
||||
|
||||
Lemma coCom i : Embeddings i.+1 << Injection i =-= Embeddings i.
|
||||
unfold Embeddings.
|
||||
rewrite -> (limitUnique (h:=(Embeddings (S i) << Injection _ : (tcone (coneN DTower i)) =-> DInf))) ; first by [].
|
||||
move => m. simpl. unfold Embeddings. rewrite -> comp_assoc. have e:= (limitCom (L DTower) (coneN DTower i.+1) m).
|
||||
simpl in e. rewrite <- e. by rewrite -> (t_nmEmbedding DTower).
|
||||
Qed.
|
||||
|
||||
Definition DCoCone : CoCone DTower := @mk_basecocone _ DInf Embeddings coCom.
|
||||
|
||||
Lemma retract_EP n : Projections n << Embeddings n =-= Id.
|
||||
unfold Projections. unfold Embeddings.
|
||||
- rewrite <- (limitCom (L DTower) (coneN DTower n) n). simpl. by rewrite -> t_nn_ID.
|
||||
Qed.
|
||||
|
||||
Lemma emp: forall i j, Projections i << Embeddings j =-= t_nm DTower j i.
|
||||
intros i j. have A:= leq_total i j. case_eq (i <= j)%N ; move => X ; rewrite X in A.
|
||||
- unfold Projections. rewrite -> (comp_eq_compat (coneCom_l (L DTower) X) (tset_refl _)).
|
||||
rewrite <- comp_assoc. rewrite -> (comp_eq_compat (tset_refl _) (retract_EP j)). by rewrite -> comp_idR.
|
||||
- simpl in A. have e:= (coconeCom_l DCoCone A). simpl in e. rewrite -> e.
|
||||
rewrite -> comp_assoc. rewrite -> (retract_EP i). by rewrite comp_idL.
|
||||
Qed.
|
||||
|
||||
Definition chainPEc (C:CoCone DTower) : cchain (cmetricMorph DInf (tcocone C)) := chainPE C (L DTower).
|
||||
|
||||
Lemma EP_id : umet_complete (chainPEc DCoCone) =-= Id.
|
||||
have Z:forall n : nat, mcone (L DTower) n =-= mcone (L DTower) n << Id by move => n ; rewrite -> comp_idR.
|
||||
rewrite <- (tset_sym (@limitUnique _ (L DTower) (L DTower) Id Z)).
|
||||
apply: limitUnique. move => n.
|
||||
have A:= (nonexp_continuous (exp_fun (ccomp _ _ _) (mcone (L DTower) n)) (chainPEc DCoCone)).
|
||||
simpl in A. rewrite -> (ccomp_eq (mcone _ n) (umet_complete (chainPEc DCoCone))) in A.
|
||||
apply: (tset_trans _ (tset_sym A)). clear A. simpl. rewrite -> (cut_complete_eq n). clear Z.
|
||||
apply tset_trans with (y:=umet_complete (const_cchain (mcone (L DTower) n:cmetricMorph _ _))) ; first by rewrite -> umet_complete_const. refine (@umet_complete_ext (cmetricMorph _ _) _ _ _).
|
||||
move => i. simpl. rewrite ccomp_eq. rewrite -> comp_assoc.
|
||||
rewrite -> (emp n (n+i)). by rewrite -> (coneCom_l (L DTower) (leq_addr i n)).
|
||||
Qed.
|
||||
|
||||
Lemma chainPE_simpl X n : chainPEc X n = mcocone X n << Projections n.
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition DCoLimit : CoLimit DTower.
|
||||
exists DCoCone (fun C => umet_complete (chainPEc C)).
|
||||
- move => C n. simpl. have ne:nonexpansive (fun x => x << Embeddings n:cmetricMorph _ _).
|
||||
move => t m e e' R. by apply comp_neq_compat. rewrite -> (cut_complete_eq n _).
|
||||
apply: (proj1 (Mrefl (_ : cmetricMorph _ _) _)). move => m.
|
||||
destruct (cumet_conv (cutn n (chainPEc C)) m) as [m' X]. specialize (X m' (leqnn _)).
|
||||
specialize (ne _ _ _ _ X). simpl in ne. refine (Mrel_trans _ ne).
|
||||
apply: (proj2 (Mrefl _ _)). rewrite <- comp_assoc. rewrite (emp _ _).
|
||||
by apply (coconeCom_l C (leq_addr m' n)).
|
||||
- move => C h X. rewrite <- (comp_idR h). rewrite <- EP_id.
|
||||
have ne:= nonexp_continuous (exp_fun (@ccomp M _ _ _) h) (chainPEc DCoCone).
|
||||
simpl in ne. rewrite -> (ccomp_eq h (umet_complete (chainPEc DCoCone))) in ne.
|
||||
rewrite -> ne. clear ne. refine (@umet_complete_ext (cmetricMorph _ _) _ _ _) => i. simpl.
|
||||
rewrite ccomp_eq. rewrite comp_assoc. by rewrite <- X.
|
||||
Defined.
|
||||
|
||||
Definition ECoCone : CoCone DTower.
|
||||
exists (ob F DInf DInf) (fun i => morph F _ _ _ _ (Projections i, Embeddings i) << Injection i).
|
||||
move => i. simpl. rewrite -> morph_comp.
|
||||
by rewrite -> (morph_eq_compat F (mconeCom (L DTower) i) (mcoconeCom DCoCone i)).
|
||||
Defined.
|
||||
|
||||
Lemma Tcomp_nonexpL (T0 T1 T2:M) (f: T0 =-> T1) : nonexpansive (fun x: T1 =-> T2 => x << f:cmetricMorph _ _).
|
||||
move => n e e' R. by apply: comp_neq_compat.
|
||||
Qed.
|
||||
|
||||
Lemma Tcomp_nonexpR (T0 T1 T2:M) (f: T2 =-> T0) : nonexpansive (fun x: T1 =-> T2 => f << x:cmetricMorph _ _).
|
||||
move => n e e' R. by apply: comp_neq_compat.
|
||||
Qed.
|
||||
|
||||
Definition FCone : Cone DTower.
|
||||
exists (ob F DInf DInf) (fun n => Projection n << morph F _ _ _ _ (Embeddings n, Projections n)).
|
||||
move => i. refine (comp_eq_compat (tset_refl _) _). unfold Projections.
|
||||
rewrite <- (morph_eq_compat F (coCom i) (mconeCom (L DTower) i)). simpl.
|
||||
by rewrite -> morph_comp.
|
||||
Defined.
|
||||
|
||||
Definition chainFPE (C:CoCone DTower) := chainPE C FCone.
|
||||
|
||||
Lemma subS_leq j i m : j.+1 - i == m.+1 -> i <= j.
|
||||
move => X. have a:=subn_gt0 i (j.+1). have x:=ltn0Sn m.
|
||||
rewrite <- (eqP X) in x. rewrite a in x. by [].
|
||||
Qed.
|
||||
|
||||
Lemma morph_tnm n m : morph F _ _ _ _ (t_nm DTower m n, t_nm DTower n m) =-= t_nm DTower n.+1 m.+1.
|
||||
have t:= (leq_total n m). case_eq (n <= m)%N ; move => l ; rewrite l in t.
|
||||
clear t. have M':exists x, m - n == x by eexists ; by []. destruct M' as [x M'].
|
||||
elim: x m n l M'.
|
||||
- move => j i l e. rewrite subn_eq0 in e.
|
||||
have A:= @anti_leq i j. rewrite l in A. rewrite e in A. specialize (A is_true_true).
|
||||
rewrite A. rewrite -> (morph_eq_compat F (t_nn_ID DTower j) (t_nn_ID DTower j)).
|
||||
rewrite -> morph_id. by rewrite -> t_nn_ID.
|
||||
- move => m IH. case.
|
||||
+ move=> i l e. rewrite leqn0 in l. have ee:=eqP l. rewrite ee.
|
||||
rewrite -> (morph_eq_compat F (t_nn_ID DTower O) (t_nn_ID DTower O)).
|
||||
rewrite -> morph_id. by rewrite -> t_nn_ID.
|
||||
+ move => j i e l. specialize (IH j i). have l':=l. have l0:=subS_leq l'. rewrite (leq_subS l0) in l.
|
||||
specialize (IH l0 l). clear l l' e.
|
||||
have e:= morph_eq_compat F (t_nmProjection2 DTower l0) (t_nmEmbedding2 DTower l0).
|
||||
rewrite <- morph_comp in e. rewrite -> e. rewrite -> IH. clear e.
|
||||
have a:= (tset_sym (t_nmEmbedding2 DTower _)).
|
||||
specialize (a i.+1 j.+1 l0). by apply a.
|
||||
simpl in t. clear l.
|
||||
have M':exists x, n - m == x by eexists ; by []. destruct M' as [x M'].
|
||||
elim: x m n t M'.
|
||||
- move => j i l e. rewrite subn_eq0 in e.
|
||||
have A:= @anti_leq i j. rewrite l in A. rewrite e in A. specialize (A is_true_true).
|
||||
rewrite A. rewrite -> (morph_eq_compat F (t_nn_ID DTower j) (t_nn_ID DTower j)).
|
||||
rewrite -> morph_id. by rewrite -> t_nn_ID.
|
||||
- move => m IH i. case.
|
||||
+ move=> l e. rewrite leqn0 in l. have ee:=eqP l. rewrite ee.
|
||||
rewrite -> (morph_eq_compat F (t_nn_ID DTower O) (t_nn_ID DTower O)).
|
||||
rewrite -> morph_id. by rewrite -> t_nn_ID.
|
||||
+ move => j e l. specialize (IH i j). have l':=l. have l0:=subS_leq l'. rewrite (leq_subS l0) in l.
|
||||
specialize (IH l0 l). clear l l' e.
|
||||
have e:= morph_eq_compat F (t_nmEmbedding2 DTower l0) (t_nmProjection2 DTower l0).
|
||||
rewrite <-morph_comp in e. rewrite -> e. rewrite -> IH.
|
||||
clear e. have a:= (tset_sym (t_nmProjection2 DTower _)).
|
||||
specialize (a j.+1 i.+1 l0). by apply a.
|
||||
Qed.
|
||||
|
||||
Lemma ECoLCom : forall (A : CoCone DTower) (n : nat),
|
||||
mcocone A n =-= umet_complete (chainFPE A) << mcocone ECoCone n.
|
||||
move => C n. simpl.
|
||||
have ne:=Tcomp_nonexpL (morph F (Diter n) DInf (Diter n) DInf (Projections n, Embeddings n) << Injection n).
|
||||
rewrite -> (cut_complete_eq n _).
|
||||
apply: (proj1 (Mrefl (_:cmetricMorph _ _) _)). move => m.
|
||||
destruct (cumet_conv (cutn n (chainFPE C)) m) as [m' X]. specialize (X m' (leqnn _)).
|
||||
specialize (ne _ _ _ _ X). simpl in ne. refine (Mrel_trans _ ne). clear ne.
|
||||
apply: (proj2 (Mrefl _ _)).
|
||||
do 2 rewrite comp_assoc. rewrite <- (comp_assoc (morph F _ _ _ _ _)). rewrite -> morph_comp.
|
||||
rewrite -> (morph_eq_compat F (emp n (n+m')) (emp (n+m') n)). rewrite morph_tnm.
|
||||
rewrite <- (comp_assoc (t_nm DTower n.+1 (n+m').+1) (Projection (n + m')) (mcocone C (n + m'))).
|
||||
rewrite <- (t_nmProjection DTower). rewrite <- comp_assoc. rewrite <- (t_nmEmbedding DTower).
|
||||
apply: (coconeCom_l C _). by apply (leq_addr _ _).
|
||||
Qed.
|
||||
|
||||
Lemma CoLimitUnique : forall (A : CoCone DTower) h,
|
||||
(forall n : nat, mcocone A n =-= h << mcocone ECoCone n) ->
|
||||
h =-= umet_complete (chainFPE A).
|
||||
move => C h X. rewrite <- (comp_idR h).
|
||||
have a:= morph_eq_compat F EP_id EP_id. rewrite -> morph_id in a. rewrite <- a. clear a.
|
||||
have b:= pair_limit (chainPEc DCoCone) (chainPEc DCoCone).
|
||||
have c:= (nonexp_continuous (morph F _ _ _ _) _).
|
||||
specialize (c _ _ _ _ (cchain_pair (chainPEc DCoCone) (chainPEc DCoCone))).
|
||||
have a:=frespect ( (morph F DInf DInf DInf DInf)) b. rewrite -> a in c. clear a b.
|
||||
rewrite -> c. clear c.
|
||||
have ne:= nonexp_continuous (exp_fun (@ccomp M _ _ _) h) ((liftc (morph F DCoCone DInf DInf DCoCone)
|
||||
(cchain_pair (chainPEc DCoCone) (chainPEc DCoCone)))). simpl in ne.
|
||||
rewrite -> (@ccomp_eq M _ _ _ h (umet_complete (liftc (morph F DCoCone DInf DInf DCoCone) _))) in ne.
|
||||
apply: (tset_trans ne). rewrite -> (cut_complete_eq 1 (chainFPE C)).
|
||||
refine (@umet_complete_ext (cmetricMorph _ _) _ _ _) => i. simpl.
|
||||
rewrite ccomp_eq. clear ne. rewrite -> morph_comp.
|
||||
rewrite <- (morph_comp F (mcone (L DTower) i) (Embeddings i) (Embeddings i) (mcone (L DTower) i)).
|
||||
rewrite comp_assoc. apply: comp_eq_compat.
|
||||
+ specialize (X i.+1). simpl in X. unfold addn. simpl. rewrite -> X. apply: comp_eq_compat ; first by [].
|
||||
rewrite -> morph_comp. apply morph_eq_compat.
|
||||
* by rewrite -> (mconeCom (L DTower) i).
|
||||
* by rewrite -> (@mcoconeCom DTower DCoCone i).
|
||||
+ apply morph_eq_compat.
|
||||
* by rewrite -> (@mcoconeCom DTower DCoCone i).
|
||||
* by rewrite -> (mconeCom (L DTower) i).
|
||||
Qed.
|
||||
|
||||
Definition ECoLimit : CoLimit DTower.
|
||||
exists ECoCone (fun C => umet_complete (chainFPE C)).
|
||||
by apply ECoLCom.
|
||||
by apply CoLimitUnique.
|
||||
Defined.
|
||||
|
||||
(*=FoldUnfoldIso *)
|
||||
Definition Fold : (ob F DInf DInf) =-> DInf := (*CLEAR*) colimitExists ECoLimit DCoCone. (*CLEARED*)
|
||||
Definition Unfold : DInf =-> (ob F DInf DInf) := (*CLEAR*) colimitExists DCoLimit ECoCone. (*CLEARED*)
|
||||
|
||||
Lemma FU_id : Fold << Unfold =-= Id.
|
||||
(*CLEAR*)
|
||||
apply tset_trans with (y:=colimitExists DCoLimit DCoLimit).
|
||||
- refine (@colimitUnique _ DCoLimit DCoLimit _ _). move => i. unfold Fold. unfold Unfold.
|
||||
simpl. rewrite <- (ccomp_eq (umet_complete (chainFPE DCoCone)) (umet_complete (chainPEc ECoCone))).
|
||||
have a:=frespect ( (ccomp _ _ _)) (pair_limit (chainFPE DCoCone) (chainPEc ECoCone)).
|
||||
rewrite -> (nonexp_continuous (ccomp _ _ _)(cchain_pair (chainFPE DCoCone) (chainPEc ECoCone)) ) in a.
|
||||
rewrite <- a. clear a.
|
||||
have ne:= nonexp_continuous (exp_fun (@ccomp M _ _ _ << <|pi2,pi1|>) (Embeddings i)) (liftc (ccomp DInf FCone DCoCone)
|
||||
(cchain_pair (chainFPE DCoCone) (chainPEc ECoCone))).
|
||||
simpl in ne.
|
||||
rewrite -> (@ccomp_eq M _ _ _ (umet_complete (liftc (ccomp DInf FCone DCoCone)
|
||||
(cchain_pair (chainFPE DCoCone) (chainPEc ECoCone)))) (Embeddings i)) in ne.
|
||||
rewrite -> ne. clear ne. rewrite -> (cut_complete_eq i _).
|
||||
apply (tset_trans (tset_sym (umet_complete_const (Embeddings i:cmetricMorph _ _)))).
|
||||
apply umet_complete_ext => j. simpl.
|
||||
rewrite ccomp_eq. simpl.
|
||||
apply: (tset_trans _ (tset_sym (comp_eq_compat (ccomp_eq _ _) (tset_refl _)))).
|
||||
do 3 rewrite comp_assoc. rewrite <- (comp_assoc _ _ (Embeddings (i + j) << Projection (i + j))).
|
||||
rewrite -> (morph_comp F (Embeddings (i + j)) (Projections (i + j)) (Projections (i + j)) (Embeddings (i + j))).
|
||||
rewrite -> (morph_eq_compat F (emp (i+j) (i+j)) (emp (i+j) (i+j))).
|
||||
rewrite -> (morph_eq_compat F (t_nn_ID DTower (i+j)) (t_nn_ID DTower (i+j))).
|
||||
rewrite morph_id. rewrite comp_idR. rewrite <- (comp_assoc (Injection (i+j))).
|
||||
rewrite retract_IP. rewrite comp_idR. rewrite <- comp_assoc. rewrite -> emp.
|
||||
apply: (coconeCom_l DCoCone). by apply leq_addr.
|
||||
- apply tset_sym. apply: (@colimitUnique _ DCoLimit DCoLimit _ _). move => n. simpl. by rewrite -> comp_idL.
|
||||
Qed.
|
||||
(*CLEARED*)
|
||||
Lemma UF_id : Unfold << Fold =-= Id.
|
||||
(*CLEAR*)
|
||||
apply tset_trans with (y:=colimitExists ECoLimit ECoLimit).
|
||||
- apply (@colimitUnique _ ECoLimit ECoLimit). move => i. unfold Fold. unfold Unfold.
|
||||
simpl. rewrite <- (ccomp_eq (umet_complete (chainPEc ECoCone)) (umet_complete (chainFPE DCoCone))).
|
||||
have a:=frespect ( (ccomp _ _ _)) (pair_limit (chainPEc ECoCone) (chainFPE DCoCone)).
|
||||
rewrite -> (nonexp_continuous (ccomp _ _ _) (cchain_pair (chainPEc ECoCone) (chainFPE DCoCone)) ) in a.
|
||||
rewrite <- a. clear a.
|
||||
have ne:= nonexp_continuous (exp_fun (@ccomp M _ _ _ << <|pi2,pi1|>) (morph F _ _ _ _ (Projections i,Embeddings i) << Injection i)) (liftc (ccomp _ _ _ )
|
||||
(cchain_pair (chainPEc ECoCone) (chainFPE DCoCone))).
|
||||
simpl in ne.
|
||||
rewrite -> (@ccomp_eq M _ _ _ (umet_complete (liftc (ccomp _ _ _)
|
||||
(cchain_pair (chainPEc ECoCone) (chainFPE DCoCone)))) (morph F _ _ _ _ (Projections i,Embeddings i) << Injection i)) in ne.
|
||||
rewrite -> ne. clear ne. rewrite -> (cut_complete_eq i _).
|
||||
apply (tset_trans (tset_sym (umet_complete_const (morph F _ _ _ _ (Projections i, Embeddings i) << Injection i:cmetricMorph _ _)))).
|
||||
apply umet_complete_ext => j. simpl.
|
||||
rewrite ccomp_eq. simpl.
|
||||
apply: (tset_trans _ (tset_sym (comp_eq_compat (ccomp_eq _ _) (tset_refl _)))).
|
||||
do 3 rewrite comp_assoc. rewrite <- (comp_assoc (Embeddings (i+j))). rewrite -> (emp (i+j) (i+j)).
|
||||
rewrite t_nn_ID. rewrite comp_idR. rewrite <- (comp_assoc (morph F _ _ _ _ (Projections i, Embeddings i))).
|
||||
rewrite -> (morph_comp F (Embeddings (i+j)) (Projections (i+j)) (Projections (i)) (Embeddings (i))).
|
||||
rewrite -> (morph_eq_compat F (emp i (i+j)) (emp (i+j) i)). rewrite morph_tnm.
|
||||
rewrite <- (comp_assoc (Injection i)). rewrite <- (t_nmEmbedding DTower i (i+j).+1).
|
||||
rewrite <- (comp_assoc _ (Projection (i+j))). rewrite <- (t_nmProjection DTower).
|
||||
rewrite <- comp_assoc. rewrite <- (@t_nmEmbedding2 DTower _ _ (leq_addr j i)). rewrite -> t_nmEmbedding.
|
||||
rewrite comp_assoc. rewrite <- morph_tnm. rewrite -> (morph_comp F (Projections (i + j)) (Embeddings (i + j))
|
||||
(t_nm DTower (i + j) i) (t_nm DTower i (i + j))).
|
||||
by rewrite <- (morph_eq_compat F (coneCom_l (L DTower) (leq_addr j i)) (coconeCom_l DCoCone (leq_addr j i))).
|
||||
- apply tset_sym. apply (@colimitUnique _ ECoLimit ECoLimit). move => n. by rewrite -> comp_idL.
|
||||
Qed.
|
||||
(*CLEARED*)
|
||||
(*=End *)
|
||||
|
||||
End RecursiveDomains.
|
||||
|
||||
End Msolution.
|
||||
|
448
papers/domains-2012/NSetoid.v
Normal file
448
papers/domains-2012/NSetoid.v
Normal file
|
@ -0,0 +1,448 @@
|
|||
(**********************************************************************************
|
||||
* NSetoid.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
Require Export Categories.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
(** printing >-> %\ensuremath{>}\ensuremath{-}\ensuremath{>}% *)
|
||||
(** printing :> %\ensuremath{:>}% *)
|
||||
(** printing Canonical %\coqdockw{Canonical}% *)
|
||||
|
||||
Open Scope S_scope.
|
||||
|
||||
Definition setoid_respect (T T':setoidType) (f: T -> T') :=
|
||||
forall x y, x =-= y -> (f x) =-= (f y).
|
||||
|
||||
Module FSet.
|
||||
|
||||
Section fset.
|
||||
Variable O1 O2 : setoidType.
|
||||
|
||||
Record mixin_of (f:O1 -> O2) := Mixin { ext :> setoid_respect f}.
|
||||
|
||||
Notation class_of := mixin_of (only parsing).
|
||||
|
||||
Section ClassDef.
|
||||
Structure type : Type := Pack {sort : O1 -> O2; _ : class_of sort; _ : O1 -> O2}.
|
||||
Local Coercion sort : type >-> Funclass.
|
||||
Definition class cT := let: Pack _ c _ := cT return class_of cT in c.
|
||||
Definition unpack K (k : forall T (c : class_of T), K T c) cT :=
|
||||
let: Pack T c _ := cT return K _ (class cT) in k _ c.
|
||||
Definition repack cT : _ -> Type -> type := let k T c p := p c in unpack k cT.
|
||||
Definition pack f (c:class_of f) := @Pack f c f.
|
||||
End ClassDef.
|
||||
End fset.
|
||||
Module Import Exports.
|
||||
Coercion sort : type >-> Funclass.
|
||||
Notation fset := type.
|
||||
Notation fsetMixin := Mixin.
|
||||
Notation fsetType := pack.
|
||||
End Exports.
|
||||
End FSet.
|
||||
Export FSet.Exports.
|
||||
|
||||
Lemma frespect (S S':setoidType) (f:fset S S') : setoid_respect f.
|
||||
case:f => f. case. move => a s. by apply a.
|
||||
Qed.
|
||||
|
||||
Hint Resolve frespect.
|
||||
|
||||
Lemma fset_setoidAxiom (T T' : setoidType) : Setoid.axiom (fun (f g:fset T T') => forall x:T, (f x) =-= (g x)).
|
||||
split ; last split.
|
||||
- by move => f.
|
||||
- move => f g h. simpl. move => X Y e. by apply (tset_trans (X e) (Y e)).
|
||||
- move => f g. simpl. move => X e. by apply (tset_sym (X e)).
|
||||
Qed.
|
||||
|
||||
Canonical Structure fset_setoidMixin (T T':setoidType) := Setoid.Mixin (fset_setoidAxiom T T').
|
||||
Canonical Structure fset_setoidType T T' := Eval hnf in SetoidType (fset_setoidMixin T T').
|
||||
|
||||
Definition setmorph_equal (A B: setoidType) : relation (fset A B) :=
|
||||
(@tset_eq (@fset_setoidType A B)).
|
||||
|
||||
Definition setmorph_equal_equivalence (A B: setoidType): Equivalence
|
||||
(@setmorph_equal A B).
|
||||
split.
|
||||
- move => x. by apply: tset_refl.
|
||||
- move => x y. by apply: tset_sym.
|
||||
- move => x y z. by apply: tset_trans.
|
||||
Defined.
|
||||
|
||||
Existing Instance setmorph_equal_equivalence.
|
||||
|
||||
Add Parametric Morphism (A B : setoidType) :
|
||||
(@FSet.sort A B) with signature (@setmorph_equal A B ==> @tset_eq A ==>
|
||||
@tset_eq B)
|
||||
as setmorph_eq_compat.
|
||||
intros f g fg x y xy. apply tset_trans with (g x) ; first by auto.
|
||||
by apply frespect.
|
||||
Qed.
|
||||
|
||||
Definition mk_fsetM (O1 O2:setoidType) (f:O1 -> O2) (m:setoid_respect f) := fsetMixin m.
|
||||
Definition mk_fset (O1 O2:setoidType) (f:O1 -> O2) (m:setoid_respect f) : fset O1 O2 :=
|
||||
Eval hnf in fsetType (mk_fsetM m).
|
||||
|
||||
Lemma ScompP S S' S'' (f:fset S' S'') (g:fset S S') : setoid_respect (T:=S) (T':=S'') (fun x : S => f (g x)).
|
||||
move => x y e. have a:=frespect g e. by apply (frespect f a).
|
||||
Qed.
|
||||
|
||||
Definition Scomp S S' S'' (f:fset S' S'') (g:fset S S') : fset S S'' :=
|
||||
Eval hnf in mk_fset (ScompP f g).
|
||||
|
||||
Lemma SidP (S:setoidType) : setoid_respect (id : S -> S).
|
||||
move => x y X. by apply X.
|
||||
Qed.
|
||||
|
||||
Definition Sid S : (fset S S) := Eval hnf in mk_fset (@SidP S).
|
||||
|
||||
Lemma setoidCatAxiom : Category.axiom Scomp Sid.
|
||||
split ; last split ; last split.
|
||||
- by move => S S' x.
|
||||
- by move => S S' x.
|
||||
- by move => S0 S1 S2 S3 f g h x.
|
||||
- move => S0 S1 S2 f f' g g' e e' x. apply tset_trans with (y:= f (g' x)).
|
||||
+ apply: (frespect f). by apply e'.
|
||||
+ by apply e.
|
||||
Qed.
|
||||
|
||||
Canonical Structure setoidCatMixin := CatMixin setoidCatAxiom.
|
||||
Canonical Structure setoidCat := Eval hnf in CatType setoidCatMixin.
|
||||
|
||||
Open Scope C_scope.
|
||||
|
||||
Lemma PropP : Setoid.axiom (fun f g : Prop => iff f g).
|
||||
split ; last split.
|
||||
- by move => x.
|
||||
- move => x y z. simpl. move => [X Y] [Z W]. by split ; auto.
|
||||
- move => x y. simpl. move => [X Y]. by split.
|
||||
Qed.
|
||||
|
||||
Canonical Structure prop_setoidMixin := Setoid.Mixin PropP.
|
||||
Canonical Structure prop_setoidType := Eval hnf in SetoidType prop_setoidMixin.
|
||||
|
||||
Notation Props := prop_setoidType.
|
||||
|
||||
Section SetoidProducts.
|
||||
Variable A B:setoidType.
|
||||
|
||||
Lemma sum_setoidAxiom : Setoid.axiom (fun p0 p1 : A + B => match p0,p1 with inl a, inl b => a =-= b | inr a, inr b => a =-= b | _,_ => False end).
|
||||
split ; last split.
|
||||
- by case.
|
||||
- case => a ; case => b ; case => c ; try done ; by apply tset_trans.
|
||||
- case => a ; case => b ; try done ; by apply tset_sym.
|
||||
Qed.
|
||||
|
||||
Canonical Structure sum_setoidMixin := Setoid.Mixin sum_setoidAxiom.
|
||||
Canonical Structure sum_setoidType := Eval hnf in SetoidType sum_setoidMixin.
|
||||
|
||||
Lemma prod_setoidAxiom : Setoid.axiom (fun p0 p1 : A * B => tset_eq (fst p0) (fst p1) /\ tset_eq (snd p0) (snd p1)).
|
||||
split ; last split.
|
||||
- case => a b ; by split.
|
||||
- case => a b ; case => c d ; case => e f ; simpl ; move => [X Y] [Z W].
|
||||
by split ; [apply (tset_trans X Z) | apply (tset_trans Y W)].
|
||||
- by case => a b ; case => c d ; simpl ; move => [X Y] ; split ; auto.
|
||||
Qed.
|
||||
|
||||
Canonical Structure prod_setoidMixin := Setoid.Mixin prod_setoidAxiom.
|
||||
Canonical Structure prod_setoidType := Eval hnf in SetoidType prod_setoidMixin.
|
||||
|
||||
Lemma Sfst_respect : setoid_respect (T:=prod_setoidType) (T':=A) (@fst _ _).
|
||||
case => a b ; case => c d ; simpl ; unlock tset_eq ; move => [X Y] ; by unlock tset_eq in X ; apply X.
|
||||
Qed.
|
||||
|
||||
Lemma Ssnd_respect : setoid_respect (T:=prod_setoidType) (T':=B) (@snd _ _).
|
||||
case => a b ; case => c d ; simpl ; unlock tset_eq ; move => [X Y] ; by unlock tset_eq in Y ; apply Y.
|
||||
Qed.
|
||||
|
||||
Definition Sfst : prod_setoidType =-> A := Eval hnf in mk_fset Sfst_respect.
|
||||
|
||||
Definition Ssnd : prod_setoidType =-> B := Eval hnf in mk_fset Ssnd_respect.
|
||||
|
||||
Lemma prod_fun_respect C (f:C =-> A) (g:C =-> B) : setoid_respect (T:=C) (T':=prod_setoidType) (fun c : C => (f c, g c)).
|
||||
move => c c' e. simpl. unlock tset_eq. by split ; [apply (frespect f e) | apply (frespect g e)].
|
||||
Qed.
|
||||
|
||||
Definition Sprod_fun C (f:C =-> A) (g:C =-> B) : C =-> prod_setoidType := Eval hnf in mk_fset (prod_fun_respect f g).
|
||||
|
||||
Lemma Sin1_respect : @setoid_respect A (sum_setoidType) (@inl _ _).
|
||||
move => a a' e. by apply e.
|
||||
Qed.
|
||||
|
||||
Lemma Sin2_respect : @setoid_respect B (sum_setoidType) (@inr _ _).
|
||||
move => a a' e. by apply e.
|
||||
Qed.
|
||||
|
||||
Definition Sin1 : A =-> sum_setoidType := Eval hnf in mk_fset Sin1_respect.
|
||||
Definition Sin2 : B =-> sum_setoidType := Eval hnf in mk_fset Sin2_respect.
|
||||
|
||||
Lemma Ssum_fun_respect C (f:A =-> C) (g:B =-> C) : @setoid_respect sum_setoidType C (fun x => match x with inl x => f x | inr x => g x end).
|
||||
case => x ; case => y => e ; try done ; by apply frespect ; apply e.
|
||||
Qed.
|
||||
|
||||
Definition Ssum_fun C (f:A =-> C) (g:B =-> C) : sum_setoidType =-> C := Eval hnf in mk_fset (Ssum_fun_respect f g).
|
||||
|
||||
|
||||
(*
|
||||
|
||||
Lemma Sprod_fun_unique C (f:C -s> A) (g:C -s> B) (h:C -s> prod_setoidType) :
|
||||
Scomp Sfst h =-= f -> Scomp Ssnd h =-= g -> h =-= Sprod_fun f g.
|
||||
move => C f g h. unlock tset_eq. move => X Y x. specialize (X x). specialize (Y x). unlock tset_eq. by split ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma Sprod_fun_fst C f g : Sfst << (@Sprod_fun C f g) =-= f.
|
||||
move => C f g. unlock tset_eq. by simpl.
|
||||
Qed.
|
||||
|
||||
Lemma Sprod_fun_snd C f g : Ssnd << (@Sprod_fun C f g) =-= g.
|
||||
move => C f g. unlock tset_eq. by simpl.
|
||||
Qed.*)
|
||||
|
||||
End SetoidProducts.
|
||||
|
||||
Lemma setoidProdCatAxiom : @CatProduct.axiom _ prod_setoidType Sfst Ssnd Sprod_fun.
|
||||
move => S0 S1 S2 f g. split ; [by split | move => m].
|
||||
move => [X Y] x. simpl. specialize (X x). by specialize (Y x).
|
||||
Qed.
|
||||
|
||||
Canonical Structure setoidProdCatMixin := prodCatMixin setoidProdCatAxiom.
|
||||
Canonical Structure setoidProdCat := Eval hnf in prodCatType (CatProduct.Class setoidProdCatMixin).
|
||||
|
||||
Lemma setoidSumCatAxiom : @CatSum.axiom _ sum_setoidType Sin1 Sin2 Ssum_fun.
|
||||
move => S0 S1 S2 f g. split ; first by split.
|
||||
move => h [X Y]. by case => x ; simpl ; [apply: X | apply: Y].
|
||||
Qed.
|
||||
|
||||
Canonical Structure setoidSumCatMixin := sumCatMixin setoidSumCatAxiom.
|
||||
Canonical Structure setoidSumCat := Eval hnf in sumCatType setoidSumCatMixin.
|
||||
|
||||
Section SetoidIProducts.
|
||||
|
||||
Variable I : Type.
|
||||
Variable P : I -> setoidType.
|
||||
|
||||
Lemma prodI_setoidP : Setoid.axiom (fun (p0 p1:forall i, P i) => forall i:I, @tset_eq _ (p0 i) (p1 i)).
|
||||
split ; last split.
|
||||
- move => X ; simpl ; by [].
|
||||
- move => X Y Z ; simpl ; move => R S i. by apply (tset_trans (R i) (S i)).
|
||||
- move => X Y ; simpl ; move => R i. by apply (tset_sym (R i)).
|
||||
Qed.
|
||||
|
||||
Canonical Structure prodI_setoidMixin := Setoid.Mixin prodI_setoidP.
|
||||
Canonical Structure prodI_setoidType := Eval hnf in SetoidType prodI_setoidMixin.
|
||||
|
||||
Lemma Sproj_respect i : setoid_respect (T:=prodI_setoidType) (T':=P i)
|
||||
(fun X : prodI_setoidType => X i).
|
||||
move => a b e. simpl. unlock tset_eq in e. by apply (e i).
|
||||
Qed.
|
||||
|
||||
Definition Sproj i : prodI_setoidType =-> P i := Eval hnf in mk_fset (Sproj_respect i).
|
||||
|
||||
Lemma SprodI_fun_respect C (f:forall i, C =-> P i) : setoid_respect (T:=C) (T':=prodI_setoidType) (fun (c : C) (i : I) => f i c).
|
||||
move => c c' e. simpl. unlock tset_eq. by move => i ; apply (frespect (f i) e).
|
||||
Qed.
|
||||
|
||||
Definition SprodI_fun C (f:forall i, C =-> P i) : C =-> prodI_setoidType := Eval hnf in mk_fset (SprodI_fun_respect f).
|
||||
|
||||
Lemma SprodI_fun_proj C (f:forall i, C =-> P i) i : Scomp (Sproj i) (SprodI_fun f) =-= f i.
|
||||
by unlock tset_eq ; simpl.
|
||||
Qed.
|
||||
|
||||
Lemma SprodI_fun_unique C (f:forall i, C =-> P i) (h:C =-> prodI_setoidType) : (forall i, Scomp (Sproj i) h =-= f i) -> h =-= SprodI_fun f.
|
||||
move => X. unlock tset_eq. move => x. unlock tset_eq. move => i. unlock tset_eq in X. apply (X i x).
|
||||
Qed.
|
||||
|
||||
End SetoidIProducts.
|
||||
|
||||
Add Parametric Morphism A B : (@pair A B)
|
||||
with signature (@tset_eq A) ==> (@tset_eq B) ==> (@tset_eq (A * B))
|
||||
as pair_eq_compat.
|
||||
by move => x y e x' y' e'.
|
||||
Qed.
|
||||
|
||||
Lemma spair_respect (A B : setoidType) (x:A) : setoid_respect (fun y:B => (x,y)).
|
||||
move => y y' e. by split.
|
||||
Qed.
|
||||
|
||||
Definition spair (A B : setoidType) (x:A) : B =-> A * B := Eval hnf in mk_fset (spair_respect x).
|
||||
|
||||
Lemma scurry_respect (A B C : setoidType) (f:A * B =-> C) : setoid_respect (fun x => f << spair _ x).
|
||||
move => x x' e y. simpl. by rewrite -> e.
|
||||
Qed.
|
||||
|
||||
Definition scurry (A B C : setoidType) (f:A * B =-> C) : A =-> (fset_setoidType B C) := Eval hnf in mk_fset (scurry_respect f).
|
||||
|
||||
Definition setoid_respect2 (S S' S'' : setoidType) (f:S -> S' -> S'') :=
|
||||
(forall s', setoid_respect (fun x => f x s')) /\ forall s, setoid_respect (f s).
|
||||
|
||||
Lemma setoid2_morphP (S S' S'' : setoidType) (f:S -> S' -> S'') (C:setoid_respect2 f) :
|
||||
setoid_respect (T:=S) (T':=S' =-> S'')
|
||||
(fun x : S => FSet.Pack (FSet.Mixin (proj2 C x)) (f x)).
|
||||
move => x y E. unfold tset_eq. simpl. move => s'. apply (proj1 C s' _ _ E).
|
||||
Qed.
|
||||
|
||||
Definition setoid2_morph (S S' S'' : setoidType) (f:S -> S' -> S'') (C:setoid_respect2 f) :
|
||||
S =-> (fset_setoidType S' S'') := Eval hnf in mk_fset (setoid2_morphP C).
|
||||
|
||||
Definition Sprod_map (A B C D:setoidType) (f:A =-> C) (g:B =-> D) := <| f << pi1, g << pi2 |>.
|
||||
|
||||
Lemma ev_respect B A : setoid_respect (fun (p:(fset_setoidType B A) * B) => fst p (snd p)).
|
||||
case => f b. case => f' b'. simpl.
|
||||
case ; simpl. move => e e'. rewrite -> e. by rewrite -> e'.
|
||||
Qed.
|
||||
|
||||
Definition Sev B A : (fset_setoidType B A) * B =-> A := Eval hnf in mk_fset (@ev_respect B A).
|
||||
|
||||
Implicit Arguments Sev [A B].
|
||||
|
||||
Lemma setoidExpAxiom : @CatExp.axiom _ fset_setoidType (@Sev) (@scurry).
|
||||
move => X Y Z h. split ; first by case.
|
||||
move => h' A x y. specialize (A (x,y)). by apply A.
|
||||
Qed.
|
||||
|
||||
Canonical Structure setoidExpMixin := expCatMixin setoidExpAxiom.
|
||||
Canonical Structure setoidExpCat := Eval hnf in expCatType setoidExpMixin.
|
||||
|
||||
Lemma const_respect (M M':setoidType) (x:M') : setoid_respect (fun _ : M => x).
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition sconst (M M':setoidType) (x:M') : M =-> M' := Eval hnf in mk_fset (const_respect x).
|
||||
|
||||
Lemma unit_setoidP : Setoid.axiom (fun _ _ : unit => True).
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Canonical Structure unit_setoidMixin := Setoid.Mixin unit_setoidP.
|
||||
Canonical Structure unit_setoidType := Eval hnf in SetoidType unit_setoidMixin.
|
||||
|
||||
Lemma setoidTerminalAxiom : CatTerminal.axiom unit_setoidType.
|
||||
move => C h h' x. case (h x). by case (h' x).
|
||||
Qed.
|
||||
|
||||
Canonical Structure setoidTerminalMixin := terminalCatMixin (fun T => sconst T tt) setoidTerminalAxiom.
|
||||
Canonical Structure setoidTerminalCat := Eval hnf in terminalCatType (CatTerminal.Class setoidTerminalMixin).
|
||||
|
||||
Lemma emp_setoid_axiom : Setoid.axiom (fun x y : Empty => True).
|
||||
split ; last split ; by move => x.
|
||||
Qed.
|
||||
|
||||
Canonical Structure emp_setoidMixin := SetoidMixin emp_setoid_axiom.
|
||||
Canonical Structure emp_setoidType := Eval hnf in SetoidType emp_setoidMixin.
|
||||
|
||||
Lemma setoidInitialAxiom : CatInitial.axiom emp_setoidType.
|
||||
move => C h h'. by case.
|
||||
Qed.
|
||||
|
||||
Lemma SZero_initial_unique D (f g : emp_setoidType =-> D) : f =-= g.
|
||||
by move => x.
|
||||
Qed.
|
||||
|
||||
Lemma SZero_respect (D:setoidType) : setoid_respect (fun (x:Empty) => match x with end : D).
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition SZero_initial D : emp_setoidType =-> D := Eval hnf in mk_fset (SZero_respect D).
|
||||
|
||||
Canonical Structure setoidInitialMixin := initialCatMixin SZero_initial setoidInitialAxiom.
|
||||
Canonical Structure setoidInitialCat := Eval hnf in initialCatType (CatInitial.Class setoidInitialMixin).
|
||||
|
||||
Section Subsetoid.
|
||||
Variable (S:setoidType) (P:S -> Prop).
|
||||
|
||||
Lemma sub_setoidP : Setoid.axiom (fun (x y:{e : S | P e}) => match x,y with exist x' _, exist y' _ => x' =-= y' end).
|
||||
split ; last split.
|
||||
- by case => x Px.
|
||||
- case => x Px. case => y Py. case => z Pz. by apply tset_trans.
|
||||
- case => x Px. case => y Py. apply tset_sym.
|
||||
Qed.
|
||||
|
||||
Canonical Structure sub_setoidMixin := Setoid.Mixin sub_setoidP.
|
||||
Canonical Structure sub_setoidType := Eval hnf in SetoidType sub_setoidMixin.
|
||||
|
||||
Lemma forget_respect : setoid_respect (fun (x:sub_setoidType) => match x with exist x' _ => x' end).
|
||||
by case => x Px ; case => y Py.
|
||||
Qed.
|
||||
|
||||
Definition Sforget : sub_setoidType =-> S := Eval hnf in mk_fset forget_respect.
|
||||
|
||||
Lemma InheretFun_respect (N : setoidType) (f:N =-> S) (c:(forall n, P (f n))) : setoid_respect (fun n => exist (fun x => P x) (f n) (c n)).
|
||||
move => n n' e. simpl. by apply (frespect f e).
|
||||
Qed.
|
||||
|
||||
Definition sInheritFun (N : setoidType) (f:N =-> S) (c:forall n, P (f n)) :
|
||||
N =-> sub_setoidType := Eval hnf in mk_fset (InheretFun_respect c).
|
||||
|
||||
Lemma Sforget_mono M (f:M =-> sub_setoidType) g : Sforget << f =-= Sforget << g -> f =-= g.
|
||||
move => C x. specialize (C x). case_eq (f x). move => fx Pf. case_eq (g x). move => gx Pg.
|
||||
move => e e'. have ee:(Sforget (f x) =-= Sforget (g x)) by []. rewrite e in ee. rewrite e' in ee. simpl in ee.
|
||||
apply ee.
|
||||
Qed.
|
||||
|
||||
End Subsetoid.
|
||||
|
||||
Section Option.
|
||||
|
||||
Variable S:setoidType.
|
||||
|
||||
Lemma option_setoid_axiom :
|
||||
Setoid.axiom (fun x y : option S => match x,y with None,None => True | Some x, Some y => x =-= y | _,_ => False end).
|
||||
split ; last split.
|
||||
- by case.
|
||||
- case ; first (move => x ; case ; [move => y ; case ; by [move => z ; apply tset_trans | by []] | by []]).
|
||||
case ; first by []. by case.
|
||||
- case ; [move => x ; case ; [move => y ; apply tset_sym | by []] | by case].
|
||||
Qed.
|
||||
|
||||
Canonical Structure option_setoidMixin := SetoidMixin option_setoid_axiom.
|
||||
Canonical Structure option_setoidType := Eval hnf in SetoidType option_setoidMixin.
|
||||
|
||||
End Option.
|
||||
|
||||
Arguments Scope fset_setoidType [S_scope S_scope].
|
||||
|
||||
Notation "'Option'" := option_setoidType : S_scope.
|
||||
|
||||
Lemma some_respect (S:setoidType) : setoid_respect (@Some S).
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition Ssome (S:setoidType) : S =-> Option S := Eval hnf in mk_fset (@some_respect S).
|
||||
|
||||
Lemma sbind_respect (S S':setoidType) (f:S =-> Option S') : setoid_respect (fun x => match x with None => None | Some x => f x end).
|
||||
case ; last by case.
|
||||
move => x. case ; last by []. move => y e. apply (frespect f). by apply e.
|
||||
Qed.
|
||||
|
||||
Definition Soptionbind (S S':setoidType) (f:S =-> Option S') : Option S =-> Option S' :=
|
||||
Eval hnf in mk_fset (sbind_respect f).
|
||||
|
||||
Lemma discrete_setoidAxiom (T:Type) : Setoid.axiom (fun x y : T => x = y).
|
||||
split ; first by [].
|
||||
split ; first by apply: trans_equal.
|
||||
by apply:sym_equal.
|
||||
Qed.
|
||||
|
||||
Definition discrete_setoidMixin T := SetoidMixin (discrete_setoidAxiom T).
|
||||
Definition discrete_setoidType T := Eval hnf in SetoidType (discrete_setoidMixin T).
|
||||
|
||||
Lemma setAxiom T : @Setoid.axiom (T -> Prop) (fun X Y => forall x, X x <-> Y x).
|
||||
split ; last split.
|
||||
- move => a. simpl. by [].
|
||||
- move => a. simpl. move => y z A B x. rewrite <- B. by rewrite <- A.
|
||||
- move => a. simpl. move => b A x. by rewrite A.
|
||||
Qed.
|
||||
|
||||
Canonical Structure setMixin T := SetoidMixin (setAxiom T).
|
||||
Canonical Structure setType T := Eval hnf in SetoidType (setMixin T).
|
||||
|
||||
Close Scope S_scope.
|
||||
Close Scope C_scope.
|
||||
|
14
papers/domains-2012/PredomAll.v
Normal file
14
papers/domains-2012/PredomAll.v
Normal file
|
@ -0,0 +1,14 @@
|
|||
(**********************************************************************************
|
||||
* PredomAll.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
Require Export PredomCore.
|
||||
Require Export PredomSum.
|
||||
Require Export PredomLift.
|
||||
Require Export PredomKleisli.
|
||||
Require Export PredomFix.
|
||||
Implicit Arguments eta [D].
|
1376
papers/domains-2012/PredomCore.v
Normal file
1376
papers/domains-2012/PredomCore.v
Normal file
File diff suppressed because it is too large
Load diff
346
papers/domains-2012/PredomFix.v
Normal file
346
papers/domains-2012/PredomFix.v
Normal file
|
@ -0,0 +1,346 @@
|
|||
(**********************************************************************************
|
||||
* PredomFix.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(*==========================================================================
|
||||
Definition of fixpoints and associated lemmas
|
||||
==========================================================================*)
|
||||
Require Import PredomCore.
|
||||
Require Import PredomLift.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
(** ** Fixpoints *)
|
||||
|
||||
Section Fixpoints.
|
||||
Variable D : cppoType.
|
||||
|
||||
Variable f : cpoCatType D D.
|
||||
|
||||
Fixpoint iter_ n : D := match n with O => PBot | S m => f (iter_ m) end.
|
||||
|
||||
Lemma iter_incr : forall n, iter_ n <= f (iter_ n).
|
||||
elim ; first by apply: leastP.
|
||||
move => n L. simpl. apply: fmonotonic. by apply L.
|
||||
Save.
|
||||
|
||||
Hint Resolve iter_incr.
|
||||
|
||||
Lemma iter_m : monotonic iter_.
|
||||
move => n n'. elim: n' n.
|
||||
- move => n. move => L. unfold Ole in L. simpl in L.
|
||||
rewrite -> (leqn0 n) in L. by rewrite (eqP L).
|
||||
- move => n IH n' L. unfold Ole in L. simpl in L. rewrite leq_eqVlt in L.
|
||||
case_eq (n' == n.+1) => E ; rewrite E in L. by rewrite (eqP E).
|
||||
specialize (IH _ L). by rewrite -> IH ; simpl.
|
||||
Qed.
|
||||
|
||||
Definition iter : natO =-> D := mk_fmono (iter_m).
|
||||
|
||||
Definition fixp : D := lub iter.
|
||||
|
||||
Lemma fixp_le : fixp <= f fixp.
|
||||
unfold fixp.
|
||||
apply Ole_trans with (lub (ocomp f iter)).
|
||||
- apply: lub_le_compat. move => x. simpl. by apply iter_incr.
|
||||
- by rewrite -> lub_comp_le.
|
||||
Save.
|
||||
Hint Resolve fixp_le.
|
||||
|
||||
Lemma fixp_eq : fixp =-= f fixp.
|
||||
apply: Ole_antisym; first by [].
|
||||
unfold fixp. rewrite {2} (lub_lift_left iter (S O)).
|
||||
rewrite -> (fcontinuous f iter). by apply: lub_le_compat => i.
|
||||
Save.
|
||||
|
||||
Lemma fixp_inv : forall g, f g <= g -> fixp <= g.
|
||||
unfold fixp; intros g l.
|
||||
apply: lub_le. elim. simpl. by apply: leastP.
|
||||
move => n L. apply: (Ole_trans _ l). apply: (fmonotonic f). by apply L.
|
||||
Save.
|
||||
|
||||
End Fixpoints.
|
||||
Hint Resolve fixp_le fixp_eq fixp_inv.
|
||||
|
||||
Definition fixp_cte (D:cppoType) : forall (d:D), fixp (const D d) =-= d.
|
||||
intros; apply fixp_eq with (f:=const D d); red; intros; auto.
|
||||
Save.
|
||||
Hint Resolve fixp_cte.
|
||||
|
||||
Add Parametric Morphism (D:cppoType) : (@fixp D)
|
||||
with signature (@Ole _ : ((D:cpoType) =-> D) -> (D =-> D) -> Prop) ++> (@Ole _)
|
||||
as fixp_le_compat.
|
||||
move => x y l. unfold fixp.
|
||||
apply: lub_le_compat. elim ; first by [].
|
||||
move => n IH. simpl. rewrite -> IH. by apply l.
|
||||
Qed.
|
||||
Hint Resolve fixp_le_compat.
|
||||
|
||||
Add Parametric Morphism (D:cppoType) : (@fixp D)
|
||||
with signature (@tset_eq _) ==> (@tset_eq D)
|
||||
as fixp_eq_compat.
|
||||
by intros x y H; apply Ole_antisym ; apply: (fixp_le_compat _) ; case: H.
|
||||
Save.
|
||||
Hint Resolve fixp_eq_compat.
|
||||
|
||||
Lemma fixp_mon (D:cppoType) : monotonic (@fixp D).
|
||||
move => x y e. by rewrite -> e.
|
||||
Qed.
|
||||
|
||||
Definition Fixp (D:cppoType) : ordCatType (D -=> D) D := Eval hnf in mk_fmono (@fixp_mon D).
|
||||
|
||||
Lemma Fixp_simpl (D:cppoType) : forall (f:D =-> D), Fixp D f = fixp f.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma iter_mon (D:cppoType) : monotonic (@iter D).
|
||||
move => x y l. elim ; first by [].
|
||||
move => n IH. simpl. rewrite -> IH. by apply l.
|
||||
Qed.
|
||||
|
||||
Definition Iter (D:cppoType) : ordCatType (D -=> D) (fmon_cpoType natO D) :=
|
||||
Eval hnf in mk_fmono (@iter_mon D).
|
||||
|
||||
Lemma IterS_simpl (D:cppoType) : forall f n, Iter D f (S n) = f (Iter _ f n).
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma iterS_simpl (D:cppoType) : forall (f:cpoCatType D D) n, iter f (S n) = f (iter f n).
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma iter_continuous (D:cppoType) :
|
||||
forall h : natO =-> D -=> D,
|
||||
iter (lub h) <= lub (Iter D << h).
|
||||
move => h. elim ; first by apply: leastP.
|
||||
move => n IH. simpl. rewrite fcont_app_eq. rewrite -> IH. rewrite <- fcont_app_eq.
|
||||
rewrite fcont_app_continuous. rewrite lub_diag. by apply lub_le_compat => i.
|
||||
Qed.
|
||||
|
||||
Hint Resolve iter_continuous.
|
||||
|
||||
Lemma iter_continuous_eq (D:cppoType) :
|
||||
forall h : natO =-> (D -=> D),
|
||||
iter (lub h) =-= lub (Iter _ << h).
|
||||
intros; apply: Ole_antisym; auto.
|
||||
exact (lub_comp_le (Iter _ (*D*)) h).
|
||||
Save.
|
||||
|
||||
|
||||
Lemma fixp_continuous (D:cppoType) : forall (h : natO =-> (D -=> D)),
|
||||
fixp (lub h) <= lub (Fixp D << h).
|
||||
move => h. unfold fixp. rewrite -> iter_continuous_eq.
|
||||
apply: lub_le => n. simpl.
|
||||
apply: lub_le => m. simpl. rewrite <- (le_lub _ m). simpl. unfold fixp. by rewrite <- (le_lub _ n).
|
||||
Save.
|
||||
Hint Resolve fixp_continuous.
|
||||
|
||||
Lemma fixp_continuous_eq (D:cppoType) : forall (h : natO =-> (D -=> D)),
|
||||
fixp (lub h) =-= lub (Fixp D << h).
|
||||
intros; apply: Ole_antisym; auto.
|
||||
by apply (lub_comp_le (Fixp D) h).
|
||||
Save.
|
||||
|
||||
Lemma Fixp_cont (D:cppoType) : continuous (@Fixp D).
|
||||
move => c.
|
||||
rewrite Fixp_simpl. by rewrite -> fixp_continuous_eq.
|
||||
Qed.
|
||||
|
||||
Definition FIXP (D:cppoType) : (D -=> D) =-> D := Eval hnf in mk_fcont (@Fixp_cont D).
|
||||
Implicit Arguments FIXP [D].
|
||||
|
||||
Lemma FIXP_simpl (D:cppoType) : forall (f:D=->D), FIXP f = fixp f.
|
||||
trivial.
|
||||
Save.
|
||||
|
||||
Lemma FIXP_le_compat (D:cppoType) : forall (f g : D =-> D),
|
||||
f <= g -> FIXP f <= FIXP g.
|
||||
move => f g l. by rewrite -> l.
|
||||
Save.
|
||||
Hint Resolve FIXP_le_compat.
|
||||
|
||||
Lemma FIXP_eq (D:cppoType) : forall (f:D=->D), FIXP f =-= f (FIXP f).
|
||||
intros; rewrite FIXP_simpl.
|
||||
by apply: (fixp_eq).
|
||||
Save.
|
||||
Hint Resolve FIXP_eq.
|
||||
|
||||
Lemma FIXP_com: forall E D (f:E =-> (D -=> D)) , FIXP << f =-= ev << <| f, FIXP << f |>.
|
||||
intros E D f. apply: fmon_eq_intro. intros e. simpl. by rewrite {1} fixp_eq.
|
||||
Qed.
|
||||
|
||||
Lemma FIXP_inv (D:cppoType) : forall (f:D=->D)(g : D), f g <= g -> FIXP f <= g.
|
||||
intros; rewrite FIXP_simpl; apply: fixp_inv; auto.
|
||||
Save.
|
||||
|
||||
(** *** Iteration of functional *)
|
||||
Lemma FIXP_comp_com (D:cppoType) : forall (f g:D=->D),
|
||||
g << f <= f << g-> FIXP g <= f (FIXP g).
|
||||
intros; apply FIXP_inv.
|
||||
apply Ole_trans with (f (g (FIXP g))).
|
||||
assert (X:=H (FIXP g)). simpl. by apply X.
|
||||
apply: fmonotonic.
|
||||
case (FIXP_eq g); trivial.
|
||||
Save.
|
||||
|
||||
Lemma FIXP_comp (D:cppoType) : forall (f g:D=->D),
|
||||
g << f <= f << g -> f (FIXP g) <= FIXP g -> FIXP (f << g) =-= FIXP g.
|
||||
intros; apply: Ole_antisym.
|
||||
- apply FIXP_inv. simpl.
|
||||
apply Ole_trans with (f (FIXP g)) ; last by []. by rewrite <- (FIXP_eq g).
|
||||
- apply: FIXP_inv.
|
||||
assert (g (f (FIXP (f << g))) <= f (g (FIXP (f << g)))).
|
||||
specialize (H (FIXP (f << g))). apply H.
|
||||
case (FIXP_eq (f<<g)); intros.
|
||||
apply Ole_trans with (2:=H3).
|
||||
apply Ole_trans with (2:=H1).
|
||||
apply: fmonotonic.
|
||||
apply FIXP_inv. simpl. apply: fmonotonic.
|
||||
apply Ole_trans with (1:=H1); auto.
|
||||
Save.
|
||||
|
||||
Fixpoint fcont_compn (D:cppoType) (f:D =->D) (n:nat) {struct n} : D =->D :=
|
||||
match n with O => f | S p => fcont_compn f p << f end.
|
||||
|
||||
Add Parametric Morphism (D1 D2 D3 : cpoType) : (@ccomp D1 D2 D3)
|
||||
with signature (@Ole (D2 -=> D3) : (D2 =-> D3) -> (D2 =-> D3) -> Prop ) ++> (@Ole (D1 -=> D2)) ++> (@Ole (D1 -=> D3))
|
||||
as fcont_comp_le_compat.
|
||||
move => f g l h k l' x. simpl. rewrite -> l. by rewrite -> l'.
|
||||
Save.
|
||||
|
||||
Lemma fcont_compn_com (D:cppoType) : forall (f:D =->D) (n:nat),
|
||||
f << (fcont_compn f n) <= fcont_compn f n << f.
|
||||
induction n; first by [].
|
||||
simpl fcont_compn. rewrite -> comp_assoc. by apply: (fcont_comp_le_compat _ (Ole_refl f)).
|
||||
Save.
|
||||
|
||||
Lemma FIXP_compn (D:cppoType) :
|
||||
forall (f:D =->D) (n:nat), FIXP (fcont_compn f n) =-= FIXP f.
|
||||
move => f. case ; first by []. simpl.
|
||||
move => n. apply: FIXP_comp ; first by apply fcont_compn_com.
|
||||
elim: n. simpl. by rewrite <- (FIXP_eq f).
|
||||
move => n IH. simpl. rewrite <- (FIXP_eq f). by apply IH.
|
||||
Save.
|
||||
|
||||
Lemma fixp_double (D:cppoType) : forall (f:D=->D), FIXP (f << f) =-= FIXP f.
|
||||
intros; exact (FIXP_compn f (S O)).
|
||||
Save.
|
||||
|
||||
|
||||
(** *** Induction principle *)
|
||||
(*=Adm *)
|
||||
Definition admissible (D:cpoType) (P:D -> Prop) :=
|
||||
forall f : natO =-> D, (forall n, P (f n)) -> P (lub f).
|
||||
Lemma fixp_ind (D:cppoType) : forall (F: D =-> D)(P:D -> Prop),
|
||||
admissible P -> P PBot -> (forall x, P x -> P (F x)) -> P (fixp F). (*CLEAR*)
|
||||
move => F P Adm B I.
|
||||
apply Adm. by elim ; simpl ; auto.
|
||||
Qed.
|
||||
(*CLEARED*)
|
||||
(*=End *)
|
||||
|
||||
Definition admissibleT (D:cpoType) (P:D -> Type) :=
|
||||
forall f : natO =-> D, (forall n, P (f n)) -> P (lub f).
|
||||
|
||||
Section SubCPO.
|
||||
Variable D E:cpoType.
|
||||
Variable P : D -> Prop.
|
||||
Variable I:admissible P.
|
||||
|
||||
Definition Subchainlub (c:natO =-> (sub_ordType P)) : sub_ordType P.
|
||||
exists (@lub D (Forgetm P << c)).
|
||||
unfold admissible in I. specialize (@I (Forgetm P << c)).
|
||||
apply I. intros i. simpl. case (c i). auto.
|
||||
Defined.
|
||||
|
||||
Lemma subCpoAxiom : CPO.axiom Subchainlub.
|
||||
intros c e n. unfold Subchainlub. split. case_eq (c n). intros x Px cn.
|
||||
refine (Ole_trans _ (@le_lub _ (Forgetm P << c) n)).
|
||||
simpl. rewrite cn. auto.
|
||||
intros C. simpl.
|
||||
case_eq e. intros dd pd de.
|
||||
refine (lub_le _). intros i. simpl. specialize (C i).
|
||||
case_eq (c i). intros d1 pd1 cn. rewrite cn in C.
|
||||
simpl in C. rewrite de in C. auto.
|
||||
Qed.
|
||||
|
||||
Canonical Structure sub_cpoMixin := CpoMixin subCpoAxiom.
|
||||
Canonical Structure sub_cpoType := Eval hnf in CpoType sub_cpoMixin.
|
||||
|
||||
Lemma InheritFun_cont (f:E =-> D) (p:forall d, P (f d)) : continuous (InheritFunm _ p).
|
||||
move => c. simpl. unfold Ole. simpl. rewrite (fcontinuous f). by apply lub_le_compat => i.
|
||||
Qed.
|
||||
|
||||
Definition InheritFun (f:E =-> D) (p:forall d, P (f d)) : E =-> sub_cpoType :=
|
||||
Eval hnf in mk_fcont (InheritFun_cont p).
|
||||
|
||||
Lemma InheritFun_simpl (f:E =-> D) (p:forall d, P (f d)) d : InheritFun p d = InheritFunm _ p d.
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Lemma Forgetm_cont : continuous (Forgetm P).
|
||||
by move => c.
|
||||
Qed.
|
||||
|
||||
Definition Forget : sub_cpoType =-> D := Eval hnf in mk_fcont Forgetm_cont.
|
||||
|
||||
Lemma forgetlub (c:natO =-> (sub_ordType P)) :
|
||||
Forget (Subchainlub c) = (@lub D (Forgetm P << c)).
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma ForgetP d : P (Forget d).
|
||||
intros. case d. auto.
|
||||
Qed.
|
||||
|
||||
End SubCPO.
|
||||
|
||||
Lemma Forget_leinj: forall (D:cpoType) (P:D -> Prop) (I:admissible P) (d:sub_cpoType I) e, Forget I d <= Forget I e -> d <= e.
|
||||
intros D P I d e. case e. clear e. intros e Pe. case d. clear d. intros d Pd.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Hint Resolve Forget_leinj.
|
||||
|
||||
Lemma Forget_inj : forall (D:cpoType) (P:D -> Prop) (I:admissible P) (d:sub_cpoType I) e, Forget I d =-= Forget I e -> d =-= e.
|
||||
intros. by split ; case: H ; case: d ; case: e.
|
||||
Qed.
|
||||
|
||||
Hint Resolve Forget_inj.
|
||||
|
||||
Lemma Forget_leinjp: forall (D E:cpoType) (P:D -> Prop) (I:admissible P) (d:E =-> sub_cpoType I) e,
|
||||
Forget I << d <= Forget I << e -> d <= e.
|
||||
intros D E P I d e C x. specialize (C x). simpl in C. unfold FCont.fmono. simpl.
|
||||
case: (e x) C => e' Pe. case: (d x) => d' Pd l. by apply l.
|
||||
Qed.
|
||||
|
||||
Lemma Forget_injp: forall (D E:cpoType) (P:D -> Prop) (I:admissible P) (d:E =-> sub_cpoType I) e,
|
||||
Forget I << d =-= Forget I << e -> d =-= e.
|
||||
intros D E P I d e C.
|
||||
apply: fmon_eq_intro. intros x. assert (CC:=fmon_eq_elim C x). simpl in CC.
|
||||
unfold FCont.fmono. simpl. clear C.
|
||||
case: (e x) CC. clear e. intros e Pe. case (d x). clear d. intros d Pd.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma InheritFun_eq_compat D E Q Qcc f g X XX : f =-= g ->
|
||||
(@InheritFun D E Q Qcc f X) =-= (@InheritFun D E Q Qcc g XX).
|
||||
intros. refine (fmon_eq_intro _). intros d. have A:=fmon_eq_elim H d. clear H. by split ; case: A.
|
||||
Qed.
|
||||
|
||||
Lemma ForgetInherit (D E:cpoType) (P:D -> Prop) PS f B : Forget PS << @InheritFun D E P PS f B =-= f.
|
||||
by refine (fmon_eq_intro _).
|
||||
Qed.
|
||||
|
||||
Lemma InheritFun_comp: forall D E F P I (f:E =-> F) X (g:D =-> E) XX,
|
||||
@InheritFun F _ P I f X << g =-= @InheritFun _ _ P I (f << g) XX.
|
||||
intros. refine (fmon_eq_intro _). intros d. simpl.
|
||||
split ; simpl ; by [].
|
||||
Qed.
|
||||
|
673
papers/domains-2012/PredomKleisli.v
Normal file
673
papers/domains-2012/PredomKleisli.v
Normal file
|
@ -0,0 +1,673 @@
|
|||
(**********************************************************************************
|
||||
* PredomKleisli.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
Require Import PredomCore.
|
||||
Require Import PredomLift.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
Definition kleislit D E (f: D -> Stream E) :=
|
||||
(*=kl *)
|
||||
cofix kl (l:Stream D) := match l with Eps l => Eps (kl l) | Val d => f d end.
|
||||
(*=End *)
|
||||
|
||||
Lemma kleisli_inv : forall D0 D1 (f: D0 -> Stream D1) (l : Stream D0), kleislit f l =
|
||||
match l with Eps l' => Eps (kleislit f l')
|
||||
| Val d => f d
|
||||
end.
|
||||
move=> D0 D1 f l.
|
||||
transitivity (match kleislit f l with Eps x => Eps x | Val d => Val d end).
|
||||
apply (DL_inv (D:=D1)).
|
||||
case l ; auto.
|
||||
move => d.
|
||||
simpl.
|
||||
symmetry.
|
||||
apply DL_inv.
|
||||
Qed.
|
||||
|
||||
(* Monad stuff *)
|
||||
Lemma left_unit : forall D0 (D1:ordType) (d:D0) (f : D0 -> Stream D1), kleislit f (Val d) =-= f d.
|
||||
intros; auto. rewrite kleisli_inv ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_Eps: forall D E (x : Stream D) (f:D -> Stream E), kleislit f (Eps x) = Eps (kleislit f x).
|
||||
intros ; rewrite kleisli_inv. auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_Val: forall D E (d:D) (f:D -> Stream E), kleislit f (Val d) = f d.
|
||||
Proof.
|
||||
intros.
|
||||
rewrite kleisli_inv.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
(* Revising kleislival to work with nats first cos that'll make the induction work better... *)
|
||||
Lemma kleislipredValVal : forall D (E:ordType) k (e:E) (M:lift_ordType D) (f : D -> lift_ordType E),
|
||||
pred_nth (kleislit f M) k = Val e -> exists d, M =-= Val d /\ f d =-= Val e.
|
||||
move => D E. elim.
|
||||
- move => e M f P.
|
||||
simpl in P. rewrite kleisli_inv in P.
|
||||
case: M P => t Eq ; first by [].
|
||||
exists t. by rewrite Eq.
|
||||
- move => n IH e M f P.
|
||||
rewrite kleisli_inv in P.
|
||||
case: M P => t Eq.
|
||||
+ specialize (IH e t f Eq). case: IH => d [A B]. exists d. split ; last by [].
|
||||
by apply (Oeq_trans (predeq (Eps t))).
|
||||
+ exists t. split ; first by [].
|
||||
apply Oeq_sym. apply (@Val_exists_pred_eq _ (f t)). by exists (S n).
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_Valeq: forall (D E:ordType) v (d:D) (f : D =-> lift_ordType E), v =-= Val d ->
|
||||
kleislit f v =-= (f d:lift_ordType E).
|
||||
move => D E v d M. case => [v1 v2].
|
||||
case: (DLle_Val_exists_pred v2) => n [dd [pd ld]].
|
||||
have DD:dd =-= d.
|
||||
- apply Ole_antisym. have xx:=DLle_pred_nth_left n v1. rewrite pd in xx.
|
||||
apply vleinj. apply xx. by apply ld.
|
||||
have X:stable M by []. specialize (X _ _ DD).
|
||||
rewrite <- X. clear ld DD X d v1 v2.
|
||||
elim: n v pd.
|
||||
- move => v. simpl. move => e. rewrite e. by rewrite kleisli_Val.
|
||||
- move => n IH. case.
|
||||
+ simpl. move => x e. rewrite -> kleisli_Eps. specialize (IH _ e). rewrite <- IH.
|
||||
by apply: (Oeq_trans (predeq _)).
|
||||
+ move => x. simpl. case => e. rewrite e. by rewrite kleisli_Val.
|
||||
Qed.
|
||||
|
||||
Hint Resolve DLle_leVal.
|
||||
|
||||
Lemma kleisli_ValVal : forall D E (M: lift_ordType D) (f : D =-> lift_ordType E) e,
|
||||
kleislit f M =-= Val e -> exists d, M =-= Val d /\ f d =-= Val e.
|
||||
move => D E M N e [kv1 kv2].
|
||||
case: (DLvalval kv2) => n [d [pd ld]].
|
||||
case: (kleislipredValVal pd) => md [vm ndm].
|
||||
exists md. split ; first by [].
|
||||
refine (Oeq_trans ndm _). move: kv1. rewrite (kleisli_Valeq N vm). rewrite -> ndm.
|
||||
move => kv1.
|
||||
split ; first by [].
|
||||
by apply: DLle_leVal.
|
||||
Qed.
|
||||
|
||||
Lemma kleislit_mono D E (X:D =-> lift_ordType E) : monotonic (kleislit X).
|
||||
move => x y L. simpl. apply: DLless_cond => e C.
|
||||
case: (kleisli_ValVal C) => dd [P Q].
|
||||
rewrite (kleisli_Valeq X P).
|
||||
rewrite -> P in L.
|
||||
destruct (DLle_Val_exists_eq L) as [y0 [yv Py]].
|
||||
rewrite (kleisli_Valeq X yv). by apply fmonotonic.
|
||||
Qed.
|
||||
|
||||
Definition kleislim D E (X:D =-> lift_ordType E) : lift_ordType D =-> lift_ordType E :=
|
||||
mk_fmono (kleislit_mono X).
|
||||
|
||||
Lemma eta_val: forall D f, @eta D f = Val f.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma eta_injp D E (f g :D =-> E) : eta << f =-= eta << g -> f =-= g.
|
||||
move => C. apply: fmon_eq_intro.
|
||||
move=> d. assert (X:=fmon_eq_elim C d).
|
||||
apply: vinj. apply X.
|
||||
Qed.
|
||||
|
||||
Lemma eta_leinjp D E (f g :D =-> E) : eta << f <= eta << g -> f <= g.
|
||||
move => C d.
|
||||
assert (X:= C d). apply vleinj. apply X.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_cont (D0 D1: cpoType) (f:D0 =-> D1 _BOT) : continuous (kleislim f).
|
||||
unfold continuous.
|
||||
intros h.
|
||||
simpl.
|
||||
apply: DLless_cond.
|
||||
intros xx C.
|
||||
destruct (kleisli_ValVal C) as [x' [P Q]].
|
||||
rewrite (kleisli_Valeq f P).
|
||||
apply Ole_trans with (y:=@lub (D1 _BOT) (kleislim f << h)) ; auto.
|
||||
clear C Q xx.
|
||||
destruct (lubval P) as [k [hk [hkv Pk]]].
|
||||
destruct (DLvalgetchain hkv) as [c Pc].
|
||||
apply Ole_trans with (y:=((lub ((f:ordCatType _ _) << c)))).
|
||||
- assert (continuous f) as cf by auto. rewrite <- cf. apply fmonotonic.
|
||||
apply vleinj. rewrite <- P. apply Ole_trans with (y:=eta (lub c)) ; last by [].
|
||||
rewrite -> (@lub_comp_eq _ _ eta c).
|
||||
rewrite -> (lub_lift_left _ k). refine (lub_le_compat _).
|
||||
intros n. by destruct (Pc n) ; auto.
|
||||
- rewrite -> (@lub_lift_left (D1 _BOT) (kleislim f << h) k).
|
||||
refine (lub_le_compat _).
|
||||
intros n.
|
||||
apply Ole_trans with (y:=kleislit f (h (k+n)%N)) ; last by auto.
|
||||
specialize (Pc n). by rewrite -> (kleisli_Valeq f Pc).
|
||||
Qed.
|
||||
|
||||
Definition kleisliX (D0 D1: cpoType) (f:D0 =-> D1 _BOT) : (D0 _BOT) =-> D1 _BOT :=
|
||||
Eval hnf in mk_fcont (kleisli_cont f).
|
||||
|
||||
Definition kleisli (D0 D1: cpoType) (f:D0 =-> D1 _BOT) : (D0 _BOT) =-> D1 _BOT := locked (kleisliX f).
|
||||
|
||||
Lemma kleisliValVal : forall D E (M: D _BOT) (f : D =-> E _BOT) e,
|
||||
kleisli f M =-= Val e -> exists d, M =-= Val d /\ f d =-= Val e.
|
||||
intros D E M N e kv. unlock kleisli in kv.
|
||||
apply (kleisli_ValVal kv).
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_simpl (D0 D1: cpoType) (f:D0 =-> D1 _BOT) v : kleisli f v = kleislit f v.
|
||||
by unlock kleisli.
|
||||
Qed.
|
||||
|
||||
Lemma kleisliVal D E d (f:D =-> E _BOT) : kleisli f (Val d) =-= f d.
|
||||
apply (@Oeq_trans _ _ (kleislit f (Val d))) ; first by unlock kleisli.
|
||||
by rewrite kleisli_Val.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_mono (D0 D1 : cpoType) : monotonic (@kleisli D0 D1).
|
||||
unfold monotonic. move => f g fgL x.
|
||||
simpl. apply: DLless_cond => e C.
|
||||
case: (kleisliValVal C) => d [P Q]. unlock kleisli. simpl.
|
||||
rewrite (kleisli_Valeq f P).
|
||||
rewrite (kleisli_Valeq g P).
|
||||
by apply fgL.
|
||||
Qed.
|
||||
|
||||
Definition Kleisli (D0 D1 : cpoType) : ordCatType (D0 -=> D1 _BOT) (D0 _BOT -=> D1 _BOT) :=
|
||||
Eval hnf in mk_fmono (@kleisli_mono D0 D1).
|
||||
|
||||
Lemma Kleisli_simpl (D0 D1 : cpoType) (f:D0 =-> D1 _BOT) : Kleisli _ _ f = kleisli f.
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Lemma Kleisli_cont (D0 D1: cpoType) : continuous (@Kleisli D0 D1).
|
||||
move => c d0. simpl.
|
||||
apply: DLless_cond.
|
||||
intros e C.
|
||||
destruct (@kleisliValVal D0 D1 d0 ((lub c)) e C) as [d [V hd]]. simpl. unlock kleisli. simpl.
|
||||
rewrite (@kleisli_Valeq _ _ _ _ ((fcont_lub c)) V). simpl.
|
||||
refine (lub_le_compat _). move => n. simpl. rewrite -> V. by rewrite kleisliVal.
|
||||
Qed.
|
||||
|
||||
Definition KLEISLI (D0 D1: cpoType) : (D0 -=> D1 _BOT) =-> D0 _BOT -=> D1 _BOT :=
|
||||
Eval hnf in mk_fcont (@Kleisli_cont D0 D1).
|
||||
|
||||
Add Parametric Morphism (D0 D1:cpoType) : (@kleisli D0 D1)
|
||||
with signature (@Ole (D0 -=> D1 _BOT)) ++> (@Ole (D0 _BOT -=> D1 _BOT))
|
||||
as kleisli_le_compat.
|
||||
move => f f' m. apply: (Ole_trans (y:=KLEISLI _ _ f)) ; first by [].
|
||||
have X:monotonic (KLEISLI D0 D1) by apply: fmonotonic.
|
||||
specialize (X _ _ m). by rewrite -> X.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D0 D1:cpoType) : (@kleisli D0 D1)
|
||||
with signature (@tset_eq _) ==> (@tset_eq _)
|
||||
as kleisli_eq_compat.
|
||||
move => f f' e. split ; [apply (kleisli_le_compat (proj1 e)) | apply (kleisli_le_compat (proj2 e))].
|
||||
Qed.
|
||||
|
||||
Implicit Arguments KLEISLI [D0 D1].
|
||||
|
||||
Lemma KLEISLI_simpl (D0 D1: cpoType) (f:D0 =-> D1 _BOT) : KLEISLI f = kleisli f.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Definition total D E (f:D -> lift_ordType E) := forall d, exists e, f d =-= eta_m e.
|
||||
|
||||
Definition DLgetelem D (x :lift_ordType D) (P:exists x', x =-= Val x') : D.
|
||||
assert (hasVal x). unfold hasVal. destruct P as [x' E].
|
||||
destruct (eqValpred E) as [n [d [X Eq]]].
|
||||
exists n. exists d. apply X.
|
||||
apply (extract (exist (@hasVal D) x H)).
|
||||
Defined.
|
||||
|
||||
Lemma total_mono D (E : ordType) (f:D =-> lift_ordType E) (Tot:total f) : monotonic (fun d => @DLgetelem E (f d) (Tot _)).
|
||||
move => x y L. apply: extractmono. simpl. assert (monotonic f) as m by auto.
|
||||
by apply m.
|
||||
Qed.
|
||||
|
||||
Definition totalLmX D (E : ordType) (f:D =-> lift_ordType E) (Tot:total f) : D =-> E :=
|
||||
Eval hnf in mk_fmono (total_mono Tot).
|
||||
|
||||
Definition totalLm D (E : ordType) (f:D =-> lift_ordType E) (Tot:total f) : D =-> E :=
|
||||
locked (totalLmX Tot).
|
||||
|
||||
Lemma total_cont (D E : cpoType) (f:D =-> E _BOT) (Tot:total f) : continuous (totalLmX Tot).
|
||||
unfold continuous. intros c. simpl. apply vleinj.
|
||||
unfold DLgetelem. rewrite <- extractworks. simpl projj.
|
||||
rewrite (fcontinuous f). rewrite -> (@lub_comp_eq _ _ eta (totalLmX Tot << c)).
|
||||
refine (lub_le_compat _) => n. simpl.
|
||||
apply (Oeq_le). simpl. unfold DLgetelem. simpl. by apply: (Oeq_trans _ (extractworks _)).
|
||||
Qed.
|
||||
|
||||
Definition totalLX (D E :cpoType) (f:D =-> E _BOT) (Tot:total f) : D =-> E :=
|
||||
Eval hnf in mk_fcont (total_cont Tot).
|
||||
|
||||
Definition totalL (D E :cpoType) (f:D =-> E _BOT) (Tot:total f) : D =-> E := locked (totalLX Tot).
|
||||
|
||||
Lemma DLgetelem_eta D (x : lift_ordType D) (P:exists x', x =-= Val x') : Val (DLgetelem P) =-= x.
|
||||
unfold DLgetelem.
|
||||
rewrite <- extractworks. simpl. auto.
|
||||
Qed.
|
||||
|
||||
Lemma totalLm_eta D (E :ordType) (f:D =-> lift_ordType E) (Tot:total f) : eta_m << totalLm Tot =-= f.
|
||||
apply: fmon_eq_intro => n. simpl. unlock totalLm. simpl. by apply DLgetelem_eta.
|
||||
Qed.
|
||||
|
||||
Lemma totalL_eta D (E :cpoType) (f:D =-> E _BOT) (Tot:total f) : eta << totalL Tot =-= f.
|
||||
assert (x:= totalLm_eta Tot).
|
||||
apply: fmon_eq_intro.
|
||||
intros d.
|
||||
assert (xx := fmon_eq_elim x d).
|
||||
rewrite <- xx. unlock totalL. by unlock totalLm.
|
||||
Qed.
|
||||
|
||||
Lemma eta_m_total D (f : D =-> lift_ordType D) : f =-= eta_m -> total f.
|
||||
move => P.
|
||||
unfold total. intros d.
|
||||
exists d.
|
||||
apply (fmon_eq_elim P d).
|
||||
Qed.
|
||||
|
||||
Lemma eta_total D (f : D =-> D _BOT) : f =-= eta -> total f.
|
||||
move => P.
|
||||
unfold total. intros d.
|
||||
exists d.
|
||||
apply (fmon_eq_elim P d).
|
||||
Qed.
|
||||
|
||||
Lemma totalLm_etap D E (f : D =-> lift_ordType E) (Tot : total f) d : eta_m (totalLm Tot d) =-= f d.
|
||||
apply (fmon_eq_elim (totalLm_eta Tot) d).
|
||||
Qed.
|
||||
|
||||
Lemma totalL_etap D E (f : D =-> E _BOT) (Tot : total f) d : eta (totalL Tot d) =-= f d.
|
||||
apply (fmon_eq_elim (totalL_eta Tot) d).
|
||||
Qed.
|
||||
|
||||
Lemma valtotalm D (E :ordType) (f:D =-> lift_ordType E) (Tot:total f) x : Val (totalLm Tot x) =-= f x.
|
||||
assert (X := fmon_eq_elim (totalLm_eta Tot) x). by apply X.
|
||||
Qed.
|
||||
|
||||
Lemma valtotal D (E :cpoType) (f:D =-> E _BOT) (Tot:total f) x : Val (totalL Tot x) =-= f x.
|
||||
assert (X := fmon_eq_elim (totalL_eta Tot) x). by apply X.
|
||||
Qed.
|
||||
|
||||
Lemma kleislit_comp: forall D E F (f : E _BOT =-> F _BOT) (g:D =-> E _BOT) d,
|
||||
(forall x xx, f x =-= Val xx -> exists y, x =-= Val y) ->
|
||||
kleislit (f << g) d =-= f (kleislit g d).
|
||||
intros D E F f g d S. split. simpl.
|
||||
apply: DLless_cond.
|
||||
|
||||
intros ff kv.
|
||||
destruct (kleisli_ValVal (f:= (f << g)) kv) as [dd [dv fdd]].
|
||||
rewrite (kleisli_Valeq ( (f << g)) dv). rewrite -> fdd.
|
||||
assert (f (g dd) =-= f (kleisli g d)). apply: fmon_stable. rewrite -> dv. by rewrite kleisliVal.
|
||||
unlock kleisli in H. simpl in H. rewrite <- H. by auto.
|
||||
|
||||
simpl. apply: DLless_cond.
|
||||
intros xx fx.
|
||||
specialize (S _ _ fx).
|
||||
destruct S as [e ke].
|
||||
destruct (kleisli_ValVal ke) as [dd [de ge]].
|
||||
rewrite (kleisli_Valeq ( (f << g)) de). apply: fmonotonic. rewrite <- ge in ke.
|
||||
by auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_comp: forall D0 D1 D2 (f:D0 =-> D1 _BOT) (g:D1 =-> D2 _BOT),
|
||||
kleisli g << kleisli f =-= kleisli (kleisli g << f).
|
||||
intros.
|
||||
refine (fmon_eq_intro _) => d. simpl. unlock {2} kleisli. simpl.
|
||||
rewrite <- (@kleislit_comp _ _ _ (kleisli g) f d) ; first by unlock kleisli.
|
||||
intros e ff.
|
||||
intros kl. destruct (kleisliValVal kl) as [dd [edd gd]].
|
||||
exists dd. by auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_leq: forall (D E F:cpoType) a b (f:D =-> E _BOT) (g: F =-> E _BOT),
|
||||
(forall aa, a =-= Val aa -> exists bb, b =-= Val bb /\ f aa <= g bb) ->
|
||||
@kleisli D E f a <= @kleisli F E g b.
|
||||
Proof.
|
||||
intros D E F a b f g V. simpl. apply: DLless_cond.
|
||||
intros xx [_ kl].
|
||||
destruct (DLle_Val_exists_pred kl) as [n [dd [pd ld]]]. unlock kleisli in pd.
|
||||
destruct (kleislipredValVal pd) as [aa [va fa]]. rewrite -> va. rewrite kleisliVal.
|
||||
specialize (V _ va).
|
||||
destruct V as [bb [vb fg]]. rewrite -> vb. rewrite kleisliVal. by apply fg.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_eq: forall (D E F:cpoType) a b (f:D =-> E _BOT) (g: F =-> E _BOT),
|
||||
(forall aa, a =-= Val aa -> exists bb, b =-= Val bb /\ f aa =-= g bb) ->
|
||||
(forall bb, b =-= Val bb -> exists aa, a =-= Val aa /\ f aa =-= g bb) ->
|
||||
@kleisli D E f a =-= @kleisli F E g b.
|
||||
Proof.
|
||||
intros D E F a b f g V1 V2.
|
||||
apply: Ole_antisym.
|
||||
apply kleisli_leq.
|
||||
intros aa va. destruct (V1 _ va) as [bb [vb fa]].
|
||||
exists bb. by auto.
|
||||
apply kleisli_leq. intros bb vb.
|
||||
destruct (V2 _ vb) as [aa [f1 f2]].
|
||||
exists aa. by auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleislim_eta_com: forall D E f, @kleislim D E f << eta_m =-= f.
|
||||
intros D E f. refine (fmon_eq_intro _) => d.
|
||||
simpl. by rewrite kleisli_Val.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_eta_com: forall D E f, @kleisli D E f << eta =-= f.
|
||||
intros D E f.
|
||||
refine (fmon_eq_intro _) => d. simpl. by apply (kleisliVal d f).
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_point_unit: forall D (d:lift_ordType D), kleislit (eta_m) d =-= d.
|
||||
intros D d.
|
||||
split. simpl. apply: DLless_cond.
|
||||
|
||||
intros dd kl.
|
||||
assert (kleislit eta_m d =-= Val dd) as kkl by auto.
|
||||
destruct (kleisli_ValVal kkl) as [d1 [P1 P2]].
|
||||
simpl in P2. rewrite -> P2 in P1.
|
||||
rewrite (kleisli_Valeq (eta_m) P1) ; auto.
|
||||
|
||||
simpl. apply: DLless_cond.
|
||||
intros dd kl.
|
||||
rewrite (kleisli_Valeq eta_m kl). auto.
|
||||
Qed.
|
||||
|
||||
Lemma kleislim_unit: forall D, kleislim (@eta_m D) =-= Id.
|
||||
intros D. refine (fmon_eq_intro _) => d. simpl.
|
||||
by rewrite kleisli_point_unit.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_unit: forall D, kleisli (@eta D) =-= Id.
|
||||
intros D. refine (fmon_eq_intro _) => d. unlock kleisli. simpl. by rewrite kleisli_point_unit.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_comp2: forall D E F (f:D =-> E) (g:E =-> F _BOT),
|
||||
kleisli (g << f) =-= kleisli g << kleisli (eta << f).
|
||||
intros D E F f g.
|
||||
refine (fmon_eq_intro _) => d.
|
||||
refine (kleisli_eq _ _).
|
||||
intros dd dv. exists (f dd). split ; last by auto.
|
||||
rewrite -> dv. by rewrite -> (kleisliVal dd (eta << f)).
|
||||
|
||||
intros ee kl.
|
||||
destruct (kleisliValVal kl) as [dd P]. exists dd.
|
||||
split. destruct P ; auto. case: P => _ P.
|
||||
simpl in P. apply: (fmon_stable g). by apply (vinj P).
|
||||
Qed.
|
||||
|
||||
Definition mu D := @kleisli (D _BOT) D Id.
|
||||
|
||||
Lemma mu_natural D E (f:D =-> E) :
|
||||
kleisli (eta << f) << mu _ =-= mu _ << kleisli (eta << (kleisli (eta << f))).
|
||||
unfold mu. rewrite <- kleisli_comp2. rewrite -> kleisli_comp.
|
||||
rewrite -> comp_idR. by rewrite comp_idL.
|
||||
Qed.
|
||||
|
||||
Lemma mumu D : mu _ << kleisli (eta << mu _) =-= mu D << mu _.
|
||||
unfold mu.
|
||||
rewrite <- kleisli_comp2. rewrite kleisli_comp. rewrite comp_idL. by rewrite comp_idR.
|
||||
Qed.
|
||||
|
||||
Lemma mukl D : mu D << kleisli (eta << eta) =-= Id.
|
||||
unfold mu.
|
||||
rewrite <- kleisli_comp2. rewrite comp_idL. by rewrite kleisli_unit.
|
||||
Qed.
|
||||
|
||||
Definition SWAP (C:prodCat) (X Y : C) : (X * Y) =-> (Y * X) := <| pi2, pi1 |>.
|
||||
Implicit Arguments SWAP [C X Y].
|
||||
|
||||
Definition Application := fun (D0 D1:cpoType) => exp_fun ((uncurry (@KLEISLI (D0 -=> D1 _BOT) D1)) <<
|
||||
<| exp_fun ((uncurry KLEISLI) << SWAP) << pi2, pi1 |>).
|
||||
|
||||
Definition Operator2 := fun A B C (F: A * B =-> C _BOT) =>
|
||||
locked ((Application _ _) << (kleisli (eta << (exp_fun F)))).
|
||||
|
||||
Definition Smash := fun A B => @Operator2 _ _ (A * B) eta.
|
||||
|
||||
Definition LStrength := fun A B => (uncurry (Smash A B)) << eta >< Id.
|
||||
Definition RStrength := fun A B => (uncurry (Smash A B)) << Id >< eta.
|
||||
|
||||
Definition KLEISLIR := fun A B C (f: A * B =-> C _BOT) => locked (kleisli f << LStrength A B).
|
||||
Definition KLEISLIL := fun A B C (f: A * B =-> C _BOT) => locked (kleisli f << RStrength A B).
|
||||
|
||||
Lemma Operator2Val: forall (D E F:cpoType) (h:D * E =-> F _BOT) d e f,
|
||||
Operator2 h d e =-= Val f -> exists d', exists e', d =-= Val d' /\ e =-= Val e' /\ h (d',e') =-= Val f.
|
||||
intros D E F h d e f. unlock Operator2.
|
||||
move => X.
|
||||
destruct (kleisliValVal X) as [f' [cc1 cc2]]. clear X.
|
||||
destruct (kleisliValVal cc2) as [e' [eeq feeq]]. clear cc2.
|
||||
destruct (kleisliValVal cc1) as [d' [deq deeq]]. clear cc1. simpl in eeq, feeq.
|
||||
exists d'. exists e'. split ; first by []. split ; first by []. rewrite <- feeq.
|
||||
by apply (fmon_eq_elim (vinj deeq) e').
|
||||
Qed.
|
||||
|
||||
Lemma Operator2_simpl: forall (E F D:cpoType) (f:E * F =-> D _BOT) v1 v2,
|
||||
Operator2 f (Val v1) (Val v2) =-= f (v1,v2).
|
||||
move => E F D f v1 v2. unlock Operator2. simpl. rewrite kleisliVal. simpl.
|
||||
rewrite kleisliVal. simpl. rewrite kleisliVal. by simpl.
|
||||
Qed.
|
||||
|
||||
Lemma strength_proj D E: kleisli (eta << pi2) << LStrength E D =-= pi2.
|
||||
apply: fmon_eq_intro. case => e d. unfold LStrength. unfold Smash. unlock Operator2.
|
||||
simpl.
|
||||
repeat (rewrite kleisliVal ; simpl). unfold id.
|
||||
apply Oeq_trans with (y:=(kleisli (eta << pi2) << kleisli (eta << PAIR D e)) d) ; first by [].
|
||||
rewrite kleisli_comp. rewrite comp_assoc. rewrite kleisli_eta_com.
|
||||
have ee:pi2 << PAIR D e =-= Id by apply: fmon_eq_intro. rewrite <- comp_assoc. rewrite ee.
|
||||
rewrite comp_idR. by rewrite kleisli_unit.
|
||||
Qed.
|
||||
|
||||
Hint Resolve tset_refl.
|
||||
|
||||
Lemma const_post_comp (X Y Z : cpoType) (f:Y =-> Z) e : f << const X e =-= const X (f e).
|
||||
by apply: fmon_eq_intro.
|
||||
Qed.
|
||||
|
||||
Lemma strength_assoc D E F :
|
||||
LStrength _ _ << <| pi1 , LStrength _ _ << pi2 |> << prod_assoc D E (F _BOT) =-=
|
||||
kleisli (eta << prod_assoc _ _ _) << LStrength _ _.
|
||||
unfold LStrength. unfold Smash. do 3 rewrite <- comp_assoc.
|
||||
rewrite (comp_assoc (prod_assoc _ _ _)). rewrite prod_map_prod_fun. do 2 rewrite comp_idL.
|
||||
refine (fmon_eq_intro _). case => p df. case: p => d e. simpl. unlock Operator2. simpl.
|
||||
repeat rewrite kleisli_Val. unfold id. rewrite (kleisliVal d). simpl.
|
||||
repeat (rewrite kleisliVal ; simpl).
|
||||
apply (@Oeq_trans _ _ ((kleisli (eta << PAIR _ d) << (kleisli (eta << PAIR _ e))) df)) ; first by [].
|
||||
apply Oeq_trans with (y:=(kleisli (eta << prod_assoc _ _ _) << kleisli (eta << PAIR _ (d,e))) df) ; last by [].
|
||||
rewrite -> kleisli_comp. rewrite comp_assoc. rewrite kleisli_eta_com.
|
||||
rewrite -> kleisli_comp. rewrite comp_assoc. rewrite kleisli_eta_com.
|
||||
apply: (fmon_eq_elim _ df). apply: kleisli_eq_compat.
|
||||
do 2 rewrite <- comp_assoc. refine (comp_eq_compat (tset_refl eta) _). by apply: fmon_eq_intro.
|
||||
Qed.
|
||||
|
||||
Lemma strength_eta D E : LStrength D E << <| pi1, eta << pi2 |> =-= eta.
|
||||
apply: fmon_eq_intro. case => d e. unfold LStrength. unfold Smash. simpl.
|
||||
by rewrite Operator2_simpl.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_eq (A A' B B' C:cpoType) : forall (f:A * B =-> C _BOT) (f': A' * B' =-> C _BOT)
|
||||
a b a' b',
|
||||
(forall aa aa', a =-= Val aa -> a' =-= Val aa' -> f (aa,b) =-= f' (aa',b')) ->
|
||||
(forall aa, a =-= Val aa -> exists aa', a' =-= Val aa') ->
|
||||
(forall aa', a' =-= Val aa' -> exists aa, a =-= Val aa) ->
|
||||
@KLEISLIL A B C f (a,b) =-= @KLEISLIL A' B' C f' (a',b').
|
||||
Proof.
|
||||
intros f f' a b a' b' V1 V2 V3. unlock KLEISLIL. simpl. unfold Smash, id.
|
||||
refine (kleisli_eq _ _).
|
||||
intros aa sa.
|
||||
destruct (Operator2Val sa) as [a1 [b1 [Pa1 [Pb1 pv]]]].
|
||||
destruct (V2 _ Pa1) as [aa' Paa].
|
||||
exists (aa', b'). specialize (V1 a1 aa' Pa1 Paa). rewrite <- V1.
|
||||
split. rewrite -> Paa.
|
||||
by rewrite Operator2_simpl. have X:=Oeq_sym (vinj pv). rewrite -> X. by rewrite (vinj Pb1).
|
||||
|
||||
intros bb ob. destruct (Operator2Val ob) as [aa' [bb' [av P]]].
|
||||
specialize (V3 _ av). destruct V3 as [aa va].
|
||||
specialize (V1 _ _ va av).
|
||||
exists (aa,b).
|
||||
split. rewrite -> va.
|
||||
rewrite Operator2_simpl. rewrite eta_val. by auto.
|
||||
refine (Oeq_trans V1 _). case: P => P1 P2. rewrite <- (vinj P2). by rewrite (vinj P1).
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIR_eq (A A' B B' C:cpoType) : forall (f:A * B =-> C _BOT) (f': A' * B' =-> C _BOT)
|
||||
a b a' b',
|
||||
(forall aa aa', b =-= Val aa -> b' =-= Val aa' -> f (a,aa) =-= f' (a',aa')) ->
|
||||
(forall aa, b =-= Val aa -> exists aa', b' =-= Val aa') ->
|
||||
(forall aa', b' =-= Val aa' -> exists aa, b =-= Val aa) ->
|
||||
@KLEISLIR A B C f (a,b) =-= @KLEISLIR A' B' C f' (a',b').
|
||||
Proof.
|
||||
intros f f' a b a' b' V1 V2 V3. unlock KLEISLIR.
|
||||
unfold LStrength. unfold Smash. simpl. unfold id.
|
||||
refine (kleisli_eq _ _).
|
||||
intros aa sa.
|
||||
destruct (Operator2Val sa) as [a1 [b1 [Pa1 [Pb1 pv]]]].
|
||||
destruct (V2 _ Pb1) as [aa' Paa].
|
||||
exists (a',aa'). specialize (V1 _ aa' Pb1 Paa). rewrite <- V1.
|
||||
split. rewrite -> Paa.
|
||||
by rewrite Operator2_simpl. rewrite <- (vinj pv). by rewrite -> (vinj Pa1).
|
||||
|
||||
intros bb ob. destruct (Operator2Val ob) as [aa' [bb' [av [bv P]]]].
|
||||
specialize (V3 _ bv). destruct V3 as [aa va].
|
||||
specialize (V1 _ _ va bv).
|
||||
exists (a,aa).
|
||||
split. rewrite -> va.
|
||||
rewrite Operator2_simpl. by auto.
|
||||
refine (Oeq_trans V1 _). rewrite <- (vinj P). by rewrite (vinj av).
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E F:cpoType) : (@KLEISLIL D E F)
|
||||
with signature (@tset_eq (D * E =-> F _BOT)) ==> (@tset_eq (D _BOT * E =-> F _BOT))
|
||||
as KLEISLIL_eq_compat.
|
||||
intros f0 f1 feq.
|
||||
apply: fmon_eq_intro. intros de. unlock KLEISLIL. simpl. by rewrite -> feq.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E F:cpoType) : (@KLEISLIR D E F)
|
||||
with signature (@tset_eq (D * E =-> F _BOT)) ==> (@tset_eq (D * (E _BOT) =-> F _BOT))
|
||||
as KLEISLIR_eq_compat.
|
||||
intros f0 f1 feq.
|
||||
apply: fmon_eq_intro. intros de. unlock KLEISLIR. simpl. by rewrite -> feq.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_comp D E F G (f:D * E =-> F _BOT) (g:G =-> D _BOT) h :
|
||||
KLEISLIL f << <| g, h|> =-= KLEISLIL (f << Id >< h ) << <|g, Id|>.
|
||||
apply: fmon_eq_intro. simpl. move => gg. unfold id. refine (KLEISLIL_eq _ _ _).
|
||||
intros d0 d1 vd0 vd1. rewrite -> vd0 in vd1. by rewrite (vinj vd1).
|
||||
intros d0 vd0. by exists d0. intros aa vv. by eexists ; apply vv.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_comp2 D E F G H (f:D * E =-> F _BOT) (g:G =-> D) (h:H =-> E) :
|
||||
KLEISLIL f << kleisli (eta << g) >< h =-= KLEISLIL (f << g >< h).
|
||||
apply: fmon_eq_intro. case => gg hh. simpl. apply: KLEISLIL_eq.
|
||||
- intros d0 d1 vd0 vd1. simpl. rewrite -> vd1 in vd0. rewrite -> kleisliVal in vd0.
|
||||
by rewrite -> (vinj vd0).
|
||||
intros d0 vd0.
|
||||
destruct (kleisliValVal vd0) as [g0 [ggl P]]. by exists g0.
|
||||
intros d0 vd0.
|
||||
exists (g d0). simpl. simpl in vd0. rewrite -> vd0. by rewrite kleisliVal.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIR_comp: forall D E F G (f:D * E =-> F _BOT) (g:G =-> E _BOT) h,
|
||||
KLEISLIR f << <|h, g|> =-= KLEISLIR (f << h >< Id) << <|Id, g|>.
|
||||
move => D E F G f g h.
|
||||
apply: fmon_eq_intro.
|
||||
intros gg.
|
||||
refine (KLEISLIR_eq _ _ _).
|
||||
intros d0 d1 vd0 vd1. simpl. unfold id. rewrite -> vd0 in vd1. by rewrite (vinj vd1).
|
||||
intros d0 vd0. by exists d0. intros aa vv. by eexists ; apply vv.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_unit: forall D E F G (f: D * E =-> F _BOT) (g:G =-> D) h,
|
||||
KLEISLIL f << <| eta << g, h |> =-= f << <| g, h |>.
|
||||
intros.
|
||||
refine (fmon_eq_intro _). intros gg. unlock KLEISLIL. simpl. unfold id. unfold Smash. rewrite Operator2_simpl.
|
||||
simpl. by rewrite kleisliVal.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIR_unit: forall D E F G (f:E * D =-> F _BOT) (g:G =-> D) h,
|
||||
KLEISLIR f << <| h, eta << g|> =-= f << <| h, g|>.
|
||||
intros.
|
||||
refine (fmon_eq_intro _). intros gg. unlock KLEISLIR. simpl. unfold Smash. rewrite Operator2_simpl.
|
||||
by rewrite kleisliVal.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIR_ValVal: forall D E F (f:D * E =-> F _BOT) e d r, KLEISLIR f (d,e) =-= Val r ->
|
||||
exists ee, e =-= Val ee /\ f (d,ee) =-= Val r.
|
||||
intros D E F f e d r. unlock KLEISLIR. simpl. unfold Smash. unfold id.
|
||||
intros klv. destruct (kleisliValVal klv) as [dd [Pdd PP]].
|
||||
destruct (Operator2Val Pdd) as [d1 [e1 [Pd1 [Pe1 XX]]]]. exists e1.
|
||||
split. by auto. rewrite <- (vinj Pd1) in XX. by rewrite -> (vinj XX).
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_ValVal: forall D E F (f: D * E =-> F _BOT) e d r, KLEISLIL f (d,e) =-= Val r ->
|
||||
exists dd, d =-= Val dd /\ f (dd,e) =-= Val r.
|
||||
intros D E F f e d r. simpl. unlock KLEISLIL. simpl. unfold Smash. unfold id.
|
||||
intros klv. destruct (kleisliValVal klv) as [dd [Pdd PP]].
|
||||
destruct (Operator2Val Pdd) as [d1 [e1 [Pd1 [Pe1 XX]]]]. exists d1. split ; first by [].
|
||||
rewrite <- (vinj Pe1) in XX. by rewrite (vinj XX).
|
||||
Qed.
|
||||
|
||||
Lemma Operator2_Valeq D E F (f:D * E =-> F _BOT) d dl e el :
|
||||
dl =-= Val d -> el =-= Val e -> Operator2 f dl el =-= f (d,e).
|
||||
move => X Y. rewrite -> X. rewrite -> Y. by rewrite Operator2_simpl.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_Valeq D E F (g:D * E =-> F _BOT) dl d e : dl =-= Val d -> KLEISLIL g (dl,e) =-= g (d,e).
|
||||
move => X. unlock KLEISLIL. simpl. unfold Smash. unfold id. rewrite -> X. rewrite Operator2_simpl.
|
||||
by rewrite kleisliVal.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_simpl D E F (g:D * E =-> F _BOT) d e : KLEISLIL g (Val d,e) =-= g (d,e).
|
||||
intros. refine (KLEISLIL_Valeq g _ (Oeq_refl _)).
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_comp3 D E F G (g : F =-> G _BOT) (f:D * E =-> F) : KLEISLIL (g << f) =-= kleisli g << KLEISLIL (eta << f).
|
||||
apply:fmon_eq_intro => x. case: x => dl e. simpl.
|
||||
split ; apply: DLless_cond.
|
||||
- intros gg kl. destruct (KLEISLIL_ValVal kl) as [d [Pdl P]]. rewrite -> (pair_eq_compat Pdl (tset_refl e)).
|
||||
do 2 rewrite KLEISLIL_simpl. simpl. by rewrite kleisliVal.
|
||||
- move => gg k. destruct (kleisliValVal k) as [ff [Pk P]].
|
||||
destruct (KLEISLIL_ValVal Pk) as [d [dlv Pd]]. rewrite -> (pair_eq_compat dlv (tset_refl e)).
|
||||
do 2 rewrite KLEISLIL_simpl. simpl. by rewrite kleisliVal.
|
||||
Qed.
|
||||
|
||||
Lemma KLEISLIL_comp4 D E F G H (f:E * G =-> H _BOT) (g:D =-> E _BOT) (h:F =-> G) :
|
||||
KLEISLIL (KLEISLIL f << g >< h) =-= KLEISLIL f << kleisli g >< h.
|
||||
intros. apply: fmon_eq_intro => x. case: x => dl df.
|
||||
split ; simpl ; apply: DLless_cond ; unfold id, Smash.
|
||||
intros hh kh. destruct (KLEISLIL_ValVal kh) as [d0 [dlv Pd]].
|
||||
destruct (KLEISLIL_ValVal Pd) as [e0 [elv Pe]]. unlock KLEISLIL.
|
||||
simpl. rewrite -> dlv. rewrite Operator2_simpl. by do 2 rewrite kleisliVal.
|
||||
|
||||
intros hh hv. destruct (KLEISLIL_ValVal hv) as [e0 [elv Pe]].
|
||||
destruct (kleisliValVal elv) as [d0 [dlv Pd]]. unlock KLEISLIL. simpl. rewrite -> dlv. rewrite Operator2_simpl.
|
||||
do 2 rewrite kleisliVal. simpl. rewrite -> Pd. by rewrite Operator2_simpl.
|
||||
Qed.
|
||||
|
||||
Definition liftCppoType (D:cpoType) := Eval hnf in CppoType (Pointed.class (@liftOrdPointedType D)).
|
||||
|
||||
Lemma strictVal: forall D E (f: D _BOT =-> E _BOT), (forall d e, f d =-= Val e -> exists dd, d =-= Val dd) -> strict f.
|
||||
intros. unfold strict. split ; last by apply: leastP.
|
||||
simpl. apply: DLless_cond.
|
||||
intros xxx Fx. specialize (H _ _ Fx). destruct H as [d incon].
|
||||
simpl in incon.
|
||||
assert (yy := eqValpred incon).
|
||||
destruct yy as [n [dd [pp _]]]. have F:False ; last by [].
|
||||
move: pp. clear. elim: n ; unfold Pointed.least ; simpl.
|
||||
- by rewrite DL_bot_eq.
|
||||
- move => n IH X. apply IH. by apply X.
|
||||
Qed.
|
||||
|
||||
Lemma kleisli_bot: forall D E (f:D =-> E _BOT), kleisli f PBot =-= PBot.
|
||||
intros.
|
||||
refine (strictVal (f:=kleisli f) _).
|
||||
intros d e.
|
||||
intros kl. destruct (kleisliValVal kl) as [dd [P _]]. by exists dd.
|
||||
Qed.
|
||||
|
899
papers/domains-2012/PredomLift.v
Normal file
899
papers/domains-2012/PredomLift.v
Normal file
|
@ -0,0 +1,899 @@
|
|||
(**********************************************************************************
|
||||
* PredomLift.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(*==========================================================================
|
||||
Lifting
|
||||
==========================================================================*)
|
||||
|
||||
Require Import PredomCore.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
(** ** Lifting *)
|
||||
|
||||
Section Lift_type.
|
||||
|
||||
(** ** Definition of underlying stream type *)
|
||||
|
||||
(** printing Eps %\eps% *)
|
||||
(** printing Val %\val% *)
|
||||
(** printing Stream %\Stream% *)
|
||||
Variable D : Type.
|
||||
|
||||
(*=Stream *)
|
||||
CoInductive Stream := Eps : Stream -> Stream | Val : D -> Stream.
|
||||
(*=End *)
|
||||
|
||||
Lemma DL_inv : forall d, d = match d with Eps x => Eps x | Val d => Val d end.
|
||||
destruct d; auto.
|
||||
Save.
|
||||
Hint Resolve DL_inv.
|
||||
|
||||
(** ** Removing Eps steps *)
|
||||
|
||||
Definition pred d : Stream := match d with Eps x => x | Val _ => d end.
|
||||
|
||||
Fixpoint pred_nth (d:Stream) (n:nat) {struct n} : Stream :=
|
||||
match n with 0 => d
|
||||
|S m => match d with Eps x => pred_nth x m
|
||||
| Val _ => d
|
||||
end
|
||||
end.
|
||||
|
||||
Lemma pred_nth_val : forall x n, pred_nth (Val x) n = Val x.
|
||||
destruct n; simpl; trivial.
|
||||
Save.
|
||||
Hint Resolve pred_nth_val.
|
||||
|
||||
Lemma pred_nth_Sn_acc : forall n d, pred_nth d (S n) = pred_nth (pred d) n.
|
||||
destruct d; simpl; auto.
|
||||
Save.
|
||||
|
||||
Lemma pred_nth_Sn : forall n d, pred_nth d (S n) = pred (pred_nth d n).
|
||||
induction n; intros; auto.
|
||||
destruct d.
|
||||
exact (IHn d).
|
||||
rewrite pred_nth_val; rewrite pred_nth_val; simpl; trivial.
|
||||
Save.
|
||||
|
||||
Lemma pred_nth_sum : forall x m n, pred_nth x (m+n) = pred_nth (pred_nth x m) n.
|
||||
induction m.
|
||||
simpl.
|
||||
intuition.
|
||||
intro n; replace (S m + n)%N with (m + S n)%N; intuition ; last by rewrite addSn ; rewrite addnS.
|
||||
rewrite IHm.
|
||||
rewrite (@pred_nth_Sn_acc n (pred_nth x m)).
|
||||
rewrite pred_nth_Sn.
|
||||
reflexivity.
|
||||
Qed.
|
||||
|
||||
End Lift_type.
|
||||
|
||||
Section Lift_ord.
|
||||
|
||||
Variable D : ordType.
|
||||
|
||||
(** ** Order *)
|
||||
(*=DLle *)
|
||||
CoInductive DLle : Stream D -> Stream D -> Prop :=
|
||||
| DLleEps : forall x y, DLle x y -> DLle (Eps x) (Eps y)
|
||||
| DLleEpsVal : forall x d, DLle x (Val d) -> DLle (Eps x) (Val d)
|
||||
| DLleVal : forall d d' n y, pred_nth y n = Val d' -> d <= d' -> DLle (Val d) y.
|
||||
(*=End *)
|
||||
|
||||
Hint Constructors DLle.
|
||||
|
||||
(*=DLle_rec *)
|
||||
Lemma DLle_rec : forall R : Stream D -> Stream D -> Prop,
|
||||
(forall x y, R (Eps x) (Eps y) -> R x y) ->
|
||||
(forall x d, R (Eps x) (Val d) -> R x (Val d)) ->
|
||||
(forall d y, R (Val d) y -> exists n, exists d', pred_nth y n = Val d' /\ d <= d')
|
||||
-> forall x y, R x y -> DLle x y.
|
||||
(*=End *)
|
||||
move => R REps REpsVal RVal; cofix; case => x.
|
||||
- case => y H.
|
||||
+ apply DLleEps. apply DLle_rec. by apply REps.
|
||||
+ by apply DLleEpsVal; apply DLle_rec; auto.
|
||||
move => y H.
|
||||
case (RVal _ _ H); move => n [d [X L]].
|
||||
by apply DLleVal with d n.
|
||||
Qed.
|
||||
|
||||
(** *** Properties of the order *)
|
||||
|
||||
Lemma DLle_refl : forall x, DLle x x.
|
||||
intros x. apply DLle_rec with (R:= fun x y => forall n, pred_nth x n = pred_nth y n) ; try (clear x).
|
||||
intros x y C n. specialize (C (S n)). simpl in C. auto.
|
||||
intros x d C n. specialize (C (S n)). simpl in C. rewrite C. case n ; auto.
|
||||
intros d y C. exists 0. exists d. simpl. specialize (C 0). simpl in C. auto.
|
||||
intros n ; auto.
|
||||
Save.
|
||||
Hint Resolve DLle_refl.
|
||||
|
||||
Lemma DLvalval : forall d x, DLle (Val d) x ->
|
||||
exists n, exists d', pred_nth x n = Val d' /\ d<=d'.
|
||||
intros d x H.
|
||||
inversion H.
|
||||
exists n; exists d'.
|
||||
intuition.
|
||||
Qed.
|
||||
|
||||
Lemma pred_nth_epsS n (x y:Stream D): Eps y = pred_nth x n -> y = pred_nth x (S n).
|
||||
elim:n x y. move => x y E. simpl in E. rewrite <- E. simpl. auto.
|
||||
move => n IH x. case x ; clear x. intros x y. intros E. simpl in E. specialize (IH _ _ E). auto.
|
||||
intros x y. simpl. intros incon. inversion incon.
|
||||
Qed.
|
||||
|
||||
Lemma pred_nth_valS n d (y:Stream D): Val d = pred_nth y n -> Val d = pred_nth y (S n).
|
||||
elim: n d y.
|
||||
- move => d y E. simpl in E. by rewrite <- E.
|
||||
- move => n IH d. case => y E.
|
||||
+ simpl in E. by rewrite (IH _ _ E).
|
||||
+ apply E.
|
||||
Qed.
|
||||
|
||||
Lemma pred_nth_comp (y:Stream D) : forall m0 n1 m1 : nat,
|
||||
m0 == (n1 + m1)%N -> pred_nth y m0 = pred_nth (pred_nth y n1) m1.
|
||||
move => m0.
|
||||
elim: m0 y.
|
||||
- move => y n1 m1 E. have X:=addn_eq0 n1 m1. rewrite <- (eqP E) in X. rewrite eq_refl in X.
|
||||
have Y:=proj1 (andP (sym_eq X)). rewrite (eqP Y). simpl. by rewrite (eqP (proj2 (andP (sym_eq X)))).
|
||||
- move => m0 IH y n1. case.
|
||||
+ rewrite addn0. case: n1 ; first by [].
|
||||
case: y ; last by [].
|
||||
simpl. move => y n E. specialize (IH y n 0). by rewrite IH ; last rewrite addn0.
|
||||
+ move => n. rewrite addnS => E. simpl. case_eq (pred_nth y n1).
|
||||
* move => y' X. case: y X ; last by case: n1 E.
|
||||
move => y X. rewrite (pred_nth_epsS (sym_eq X)). simpl. by apply IH.
|
||||
* move => y' X. case: y X => y.
|
||||
- case: n1 E ; first by []. move => n1 E. simpl.
|
||||
specialize (IH y n1). move => X. rewrite X in IH. specialize (IH n.+1). rewrite addnS in IH. specialize (IH E).
|
||||
by apply IH.
|
||||
- move => X. have EE:= (pred_nth_val y n1). rewrite EE in X. by apply X.
|
||||
Qed.
|
||||
|
||||
Lemma DLle_pred_nth x y n : DLle x y -> DLle (pred_nth x n) (pred_nth y n).
|
||||
elim: n x y ; first by [].
|
||||
move => n IH x y C. case: x y / C.
|
||||
- move => x y C. simpl. apply IH. by apply C.
|
||||
- move => x d C. simpl. specialize (IH _ _ C). rewrite pred_nth_val in IH. by apply IH.
|
||||
- move => d d' m y E L. specialize (IH (Val d)).
|
||||
case: y E.
|
||||
+ move => y. specialize (IH y). move => X. simpl. rewrite (pred_nth_val) in IH. apply IH.
|
||||
case: m X ; first by [].
|
||||
move => m X. simpl in X. by apply:(DLleVal X).
|
||||
+ move => y. rewrite pred_nth_val. case. move => E. simpl. rewrite E.
|
||||
by apply DLleVal with d' 0 ; last by [].
|
||||
Qed.
|
||||
|
||||
Lemma DLleEps_right : forall x y, DLle x y -> DLle x (Eps y).
|
||||
intros x y C. apply DLle_rec with (R:= fun x y => forall n, DLle (pred_nth x n) (pred_nth y (S n))).
|
||||
clear x y C. intros x y C n. specialize (C (S n)). auto.
|
||||
clear x y C. intros x y C n. specialize (C (S n)). auto.
|
||||
clear x y C. intros d y C. specialize (C 0).
|
||||
assert (DLle (Val d) (pred_nth y 1)) as CC by auto. clear C.
|
||||
destruct (DLvalval CC) as [n [dd [X Y]]]. exists (S n). exists dd. split ; auto.
|
||||
generalize X. clear X. destruct y. simpl. auto. simpl. case n ; auto.
|
||||
intros n. simpl.
|
||||
apply DLle_pred_nth ; auto.
|
||||
Qed.
|
||||
Hint Resolve DLleEps_right.
|
||||
|
||||
Lemma DLleEps_left : forall x y, DLle x y -> DLle (Eps x) y.
|
||||
intros x y C. apply DLle_rec with (R:= fun x y => forall n, DLle (pred_nth x (S n)) (pred_nth y n)).
|
||||
clear x y C. intros x y C n. specialize (C (S n)). auto.
|
||||
clear x y C. intros x y C n. specialize (C (S n)). destruct n ; auto.
|
||||
clear x y C. intros d y C. specialize (C 0).
|
||||
simpl in C.
|
||||
destruct (DLvalval C) as [n [dd [X Y]]]. exists n. exists dd. split ; auto.
|
||||
intros n. simpl. apply DLle_pred_nth ; auto.
|
||||
Qed.
|
||||
Hint Resolve DLleEps_left.
|
||||
|
||||
Lemma DLle_pred_left : forall x y, DLle x y -> DLle (pred x) y.
|
||||
destruct x; destruct y; simpl; intros; trivial.
|
||||
inversion H; auto.
|
||||
inversion H; trivial.
|
||||
Save.
|
||||
|
||||
Lemma DLle_pred_right : forall x y, DLle x y -> DLle x (pred y).
|
||||
destruct x; destruct y; simpl; intros; trivial.
|
||||
inversion H; auto.
|
||||
inversion H; trivial.
|
||||
destruct n; simpl in H1.
|
||||
discriminate H1.
|
||||
apply DLleVal with d' n; auto.
|
||||
Save.
|
||||
|
||||
Hint Resolve DLle_pred_left DLle_pred_right.
|
||||
|
||||
Lemma DLle_pred : forall x y, DLle x y -> DLle (pred x) (pred y).
|
||||
auto.
|
||||
Save.
|
||||
|
||||
Hint Resolve DLle_pred.
|
||||
|
||||
Lemma DLle_pred_nth_left : forall n x y, DLle x y -> DLle (pred_nth x n) y.
|
||||
induction n; intros.
|
||||
simpl; auto.
|
||||
rewrite pred_nth_Sn; auto.
|
||||
Save.
|
||||
|
||||
Lemma DLle_pred_nth_right : forall n x y,
|
||||
DLle x y -> DLle x (pred_nth y n).
|
||||
induction n; intros.
|
||||
simpl; auto.
|
||||
rewrite pred_nth_Sn; auto.
|
||||
Save.
|
||||
|
||||
Hint Resolve DLle_pred_nth_left DLle_pred_nth_right.
|
||||
|
||||
|
||||
(* should be called leq *)
|
||||
Lemma DLleVal_leq : forall x y, DLle (Val x) (Val y) -> x <= y.
|
||||
intros x y H; inversion H.
|
||||
destruct n; simpl in H1; injection H1;intro; subst y; auto.
|
||||
Save.
|
||||
Hint Immediate DLleVal_leq.
|
||||
|
||||
(* New *)
|
||||
Lemma DLle_leVal : forall x y, x<=y -> DLle (Val x) (Val y).
|
||||
intros x y H.
|
||||
apply (DLle_rec (R := fun x' y' => x'=Val x /\ y'=Val y)).
|
||||
intros.
|
||||
destruct H0; discriminate.
|
||||
intros.
|
||||
destruct H0; discriminate.
|
||||
intros.
|
||||
destruct H0.
|
||||
subst y0.
|
||||
exists 0.
|
||||
exists y.
|
||||
simpl.
|
||||
split; trivial.
|
||||
injection H0.
|
||||
intro; subst d; auto.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma DLnvalval : forall n y d z, pred_nth y n = Val d -> DLle y z ->
|
||||
exists m, exists d', pred_nth z m = Val d' /\ d <= d'.
|
||||
elim.
|
||||
- move => y d z E L. simpl in E. case: L E ; [by [] | by [] | idtac].
|
||||
move => d1 d2 n y' E L. case => E'. rewrite E' in L. clear E' d1.
|
||||
exists n. exists d2. by split.
|
||||
- move => n IH y d z E L. specialize (IH (pred y) d z). rewrite pred_nth_Sn_acc in E.
|
||||
specialize (IH E). apply IH. by apply DLle_pred_left.
|
||||
Qed.
|
||||
|
||||
Lemma DLleEpsn n x z xx : pred_nth x n = Val xx -> DLle (Val xx) z -> DLle x z.
|
||||
elim:n x z xx => x.
|
||||
- move => z xx. simpl => E. by rewrite E.
|
||||
- move => IH. case.
|
||||
+ move => d z xx. simpl. move => C DD.
|
||||
specialize (IH d z xx C DD). by apply: (DLle_pred_right (DLleEps _)).
|
||||
+ move => d z xx. simpl. case => e. by rewrite e.
|
||||
Qed.
|
||||
|
||||
Lemma DLle_trans : forall x y z, DLle x y -> DLle y z -> DLle x z.
|
||||
move => x y z D0 D1.
|
||||
apply DLle_rec with (R:=fun x z => (forall xx n, pred_nth x n = Val xx -> DLle x z)).
|
||||
- clear x y z D0 D1. move => x y C xx n X.
|
||||
specialize (C xx (S n) X). by apply (DLle_pred C).
|
||||
- clear x y z D0 D1. move => x y C xx n X. specialize (C _ (S n) X).
|
||||
by apply (DLle_pred_left C).
|
||||
- move => d yy C. specialize (C d 0).
|
||||
case (fun Z => DLvalval (C Z)) ; first by [].
|
||||
move => n [dd [ZZ Z]]. exists n. exists dd. by split.
|
||||
- move => xx n C. case (DLnvalval C D0) => m [yy [P Q]].
|
||||
case (DLnvalval P D1) => k [zz [Pz Qz]].
|
||||
apply (DLleEpsn C). apply (DLleVal Pz). by apply (Ole_trans Q Qz).
|
||||
Qed.
|
||||
|
||||
(** *** Declaration of the ordered set *)
|
||||
|
||||
Lemma LiftOrdAxiom : PreOrd.axiom DLle.
|
||||
move => x. split ; first by [].
|
||||
move => y z. by apply DLle_trans.
|
||||
Qed.
|
||||
|
||||
Canonical Structure lift_ordMixin := OrdMixin LiftOrdAxiom.
|
||||
Canonical Structure lift_ordType := Eval hnf in OrdType lift_ordMixin.
|
||||
|
||||
Lemma ordSetoidAxiom (X:ordType) : Setoid.axiom (@tset_eq X).
|
||||
split ; first by [].
|
||||
split ; last by apply: tset_sym.
|
||||
by apply: tset_trans.
|
||||
Qed.
|
||||
|
||||
Canonical Structure ordSetoidMixin D := SetoidMixin (ordSetoidAxiom D).
|
||||
Canonical Structure ordSetoidType D := Eval hnf in SetoidType (ordSetoidMixin D).
|
||||
|
||||
Lemma lift_setoidAxiom : @Setoid.axiom (Stream D) (@tset_eq lift_ordType).
|
||||
split ; last split ; first by [].
|
||||
move => x y z. by apply tset_trans.
|
||||
move => x y ; by apply tset_sym.
|
||||
Qed.
|
||||
|
||||
Canonical Structure lift_setoidMixin := SetoidMixin lift_setoidAxiom.
|
||||
Canonical Structure lift_setoidType := Eval hnf in SetoidType lift_setoidMixin.
|
||||
|
||||
(** ** Definition of the cpo structure *)
|
||||
|
||||
Lemma eq_Eps : forall x, x =-= Eps x.
|
||||
intros; apply: Ole_antisym; repeat red; auto.
|
||||
Save.
|
||||
Hint Resolve eq_Eps.
|
||||
|
||||
(** *** Bottom is given by an infinite chain of Eps *)
|
||||
(** printing DL_bot %\DLbot% *)
|
||||
CoFixpoint DL_bot : lift_ordType := Eps DL_bot.
|
||||
|
||||
Lemma DL_bot_eq : DL_bot = Eps DL_bot.
|
||||
transitivity match DL_bot with Eps x => Eps x | Val d => Val d end ; auto.
|
||||
destruct DL_bot ; auto.
|
||||
Save.
|
||||
|
||||
Lemma DLless_cond : forall x y, (forall xx, x =-= Val xx -> x <= y) -> DLle x y.
|
||||
move => x y P. apply DLle_rec with (R:=fun x y => forall xx, x =-= Val xx -> x <= y).
|
||||
- move => d0 d1 IH d00 dv.
|
||||
rewrite -> (eq_Eps d0) in dv. rewrite -> (eq_Eps d0). specialize (IH _ dv).
|
||||
by rewrite -> (eq_Eps d1).
|
||||
- move => d0 d1 IH dd dv. rewrite -> (eq_Eps d0) in dv. specialize (IH _ dv).
|
||||
by rewrite -> (eq_Eps d0).
|
||||
- move => d0 d1 IH. specialize (IH _ (Oeq_refl _)). by apply (DLvalval IH).
|
||||
- by apply P.
|
||||
Qed.
|
||||
|
||||
Lemma DL_bot_least : forall x, DL_bot <= x.
|
||||
move => x. apply DLless_cond => xx [_ C].
|
||||
case (DLvalval C) => n [dd [P Q]].
|
||||
have F:False ; last by [].
|
||||
elim: n P ; first by rewrite DL_bot_eq.
|
||||
move => n IH. rewrite DL_bot_eq. by apply IH.
|
||||
Qed.
|
||||
|
||||
(** *** More properties of elements in the lifted domain *)
|
||||
|
||||
(* This is essentially the same, when fixed up, as DLvalval that I proved above
|
||||
I've repeated it here, though
|
||||
*)
|
||||
Lemma DLle_Val_exists_pred :
|
||||
forall x d, Val d <= x -> exists k, exists d', pred_nth x k = Val d'
|
||||
/\ d <= d'.
|
||||
intros x d H; inversion H; eauto.
|
||||
Save.
|
||||
|
||||
Lemma Val_exists_pred_le :
|
||||
forall x d, (exists k, pred_nth x k = Val d) -> Val d <= x.
|
||||
destruct 1; intros.
|
||||
apply DLleVal with d x0; trivial.
|
||||
Save.
|
||||
Hint Immediate DLle_Val_exists_pred Val_exists_pred_le.
|
||||
|
||||
Lemma Val_exists_pred_eq :
|
||||
forall x d, (exists k, pred_nth x k = Val d) -> (Val d:lift_ordType) =-= x.
|
||||
move => x d X. split. simpl. case: X => k e. by apply: (DLleVal e).
|
||||
case: X => n X. rewrite <- X. by apply DLle_pred_nth_right.
|
||||
Save.
|
||||
|
||||
(** *** Construction of least upper bounds *)
|
||||
|
||||
Definition isEps (x:Stream D) := match x with Eps _ => True | _ => False end.
|
||||
|
||||
Lemma isEps_Eps : forall x, isEps (Eps x).
|
||||
repeat red; auto.
|
||||
Save.
|
||||
|
||||
Lemma not_isEpsVal : forall d, ~ (isEps (Val d)).
|
||||
repeat red; auto.
|
||||
Save.
|
||||
Hint Resolve isEps_Eps not_isEpsVal.
|
||||
|
||||
Lemma isEpsEps (x x': Stream D) : x = Eps x' -> isEps x.
|
||||
move => E. by rewrite E.
|
||||
Qed.
|
||||
|
||||
Definition isEps_dec (x: Stream D) : {d:D|x=Val d}+{isEps x} :=
|
||||
match x as x0 return x = x0 -> {d:D|x=Val d}+{isEps x} with
|
||||
| Eps x' => fun E => inright _ (isEpsEps E)
|
||||
| Val d => fun E => inleft _ (exist _ d E)
|
||||
end (refl_equal x).
|
||||
|
||||
|
||||
(* Slightly more radical rewrite, trying to use equality *)
|
||||
|
||||
Lemma allvalsfromhere : forall (c:natO =-> lift_ordType) n d i,
|
||||
c n =-= Val d -> exists d', c (n+i)%N =-= Val d' /\ d <= d'.
|
||||
move => c n d i [Hnv1 Hnv2].
|
||||
have X:Val d <= c (n+i)%N. apply Ole_trans with (c n) ; first by []. apply fmonotonic. by apply: leq_addr.
|
||||
case (DLle_Val_exists_pred X) => k [d' [P Q]].
|
||||
exists d'. split ; last by [].
|
||||
apply Oeq_sym. apply Val_exists_pred_eq. by exists k.
|
||||
Qed.
|
||||
|
||||
Definition hasVal (x:Stream D) := exists j, exists d, pred_nth x j = Val d.
|
||||
Definition valuable := {x | hasVal x}.
|
||||
|
||||
Definition projj (x:valuable) : lift_ordType :=
|
||||
match x with exist x X => x end.
|
||||
|
||||
(* Redoing extract using constructive epsilon *)
|
||||
Require Import ConstructiveEpsilon.
|
||||
|
||||
Lemma hasValEps x x' : x = Eps x' -> hasVal x -> hasVal x'.
|
||||
move => e. rewrite e. clear e x.
|
||||
case => n. case => d e. case: n e ; first by []. move => n. simpl => e. exists n. exists d. by apply e.
|
||||
Qed.
|
||||
|
||||
Definition inc x' (X:{kd : nat * D | let (k, d) := kd in pred_nth x' k = Val d}) :
|
||||
{kd : nat * D | let (k, d) := kd in pred_nth (Eps x') k = Val d} :=
|
||||
match X with exist (k,d) X0 =>
|
||||
exist (fun kd : nat * D => let (k, d) := kd in pred_nth (Eps x') k = Val d) (k.+1,d) X0
|
||||
end.
|
||||
|
||||
Lemma pred_nth_notEps x' : ~ (exists d : D, pred_nth (Eps x') 0 = Val d).
|
||||
by case.
|
||||
Qed.
|
||||
|
||||
Lemma pred_nthvalval d' n : (exists d : D, pred_nth (Val d') n = Val d).
|
||||
exists d'. by case: n.
|
||||
Qed.
|
||||
|
||||
Fixpoint hasVal_dec x n : { (exists d : D, pred_nth x n = Val d) } + {~ (exists d : D, pred_nth x n = Val d)} :=
|
||||
match x,n return { (exists d : D, pred_nth x n = Val d) } + {~ (exists d : D, pred_nth x n = Val d)} with
|
||||
| Eps x', S n => hasVal_dec x' n
|
||||
| Eps x', O => right (exists d : D, pred_nth (Eps x') O = Val d) (@pred_nth_notEps x')
|
||||
| Val d, n' => left (~ (exists d' : D, pred_nth (Val d) n = Val d')) (pred_nthvalval d n)
|
||||
end.
|
||||
|
||||
Fixpoint getVal x n :=
|
||||
match x as x0, n as n0 return (exists d, pred_nth x0 n0 = Val d) -> D with
|
||||
| Eps x', S n' => getVal x' n'
|
||||
| Eps x', O => fun P => match pred_nth_notEps P with end
|
||||
| Val d, n => fun _ => d
|
||||
end.
|
||||
|
||||
Lemma getValVal n x (P:exists d, pred_nth x n = Val d) : pred_nth x n = Val (@getVal x n P).
|
||||
elim: n x P.
|
||||
- by case ; first by move => x [d F].
|
||||
- move => n IH. by case; first by apply IH.
|
||||
Qed.
|
||||
|
||||
Definition findindexandval (x:valuable) : {kd | match kd with (k,d) => (pred_nth (projj x) k) = Val d end} :=
|
||||
match x as x0 return {kd | match kd with (k,d) => (pred_nth (projj x0) k) = Val d end} with
|
||||
| exist x' P =>
|
||||
match @constructive_indefinite_description_nat (fun n => exists d:D, pred_nth x' n = Val d) (hasVal_dec x') P
|
||||
with exist n Px => exist (fun kd => match kd with (k,d) => (pred_nth x' k) = Val d end) (n,getVal Px) (getValVal Px)
|
||||
end
|
||||
end.
|
||||
|
||||
Definition extract (x:valuable) : D :=
|
||||
match (findindexandval x) with exist (n,d) _ => d end.
|
||||
|
||||
Lemma extractworks x : projj x =-= Val (extract x).
|
||||
unfold extract. case (findindexandval x).
|
||||
case => n d. move => X. apply Oeq_sym. apply Val_exists_pred_eq. exists n. by apply X.
|
||||
Qed.
|
||||
|
||||
Lemma vinj : forall d d', (Val d:lift_ordType) =-= Val d' -> d =-= d'.
|
||||
intros d d' (H1,H2). by split ; apply DLleVal_leq.
|
||||
Qed.
|
||||
|
||||
Lemma vleinj : forall d d', Val d <= Val d' -> d <= d'.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma extractmono : forall (x y : valuable), (projj x) <= (projj y) -> extract x <= extract y.
|
||||
intros x y H.
|
||||
apply vleinj.
|
||||
rewrite <- (extractworks x).
|
||||
rewrite <- (extractworks y).
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
(* This is the simpler one that just takes an equality *)
|
||||
|
||||
Lemma pred_nth_eq k x : pred_nth x k =-= x.
|
||||
elim: k x; first by [].
|
||||
move => n IH. case ; last by [].
|
||||
simpl. move => x. rewrite -> (IH x). split ; first by apply DLleEps_right. by apply DLleEps_left.
|
||||
Qed.
|
||||
|
||||
Lemma eqValpred : forall x d, x =-= Val d -> exists k, exists d', pred_nth x k = Val d' /\ d=-=d'.
|
||||
move => x d [H1 H2].
|
||||
case (DLle_Val_exists_pred H2) => k [d' [H L]].
|
||||
exists k; exists d';split ; first by [].
|
||||
split ; first by [].
|
||||
apply vleinj. rewrite <- H. clear L d' H.
|
||||
rewrite <- H1. by rewrite -> pred_nth_eq.
|
||||
Qed.
|
||||
|
||||
Lemma hasValShift (c:natO -=> lift_ordType) k d (Hck:c k =-= Val d) n : hasVal (c (n + k)%N).
|
||||
rewrite addnC. case: (allvalsfromhere n Hck) => d' [A L].
|
||||
case: (eqValpred A) => i [dd [XX _]]. exists i. by exists dd.
|
||||
Qed.
|
||||
|
||||
Definition makechain (c:natO -=> lift_ordType) k d (Hck:c k =-= Val d) : nat -> D :=
|
||||
fun n => @extract (exist hasVal (c (n+k)%N) (hasValShift Hck n)).
|
||||
|
||||
Lemma makechan_mono (c:natO -=> lift_ordType) k d (Hck:c k =-= Val d) : monotonic (makechain Hck).
|
||||
move => n n' L. apply extractmono. simpl. apply fmonotonic. by apply: (leq_add L).
|
||||
Qed.
|
||||
|
||||
Definition makechainm (c:natO -=> lift_ordType) k d (Hck:c k =-= Val d) : natO =-> D :=
|
||||
Eval hnf in mk_fmono (makechan_mono Hck).
|
||||
|
||||
Lemma pred_mono (c:natO =-> lift_ordType) : monotonic (fun n => pred (c n)).
|
||||
move => x y H. apply: DLle_pred. by apply (fmonotonic c).
|
||||
Qed.
|
||||
|
||||
Definition cpred (c:natO =-> lift_ordType) : natO =-> lift_ordType := Eval hnf in mk_fmono (pred_mono c).
|
||||
|
||||
Lemma fValIdx : forall (c:natO =-> lift_ordType) (n : nat),
|
||||
{dk: (D*nat)%type | match dk with (d,k) => k<n /\ c k =-= Val d end} + {(forall k, k<n -> isEps (c k))}.
|
||||
move => c. elim ; first by right.
|
||||
move => n. case.
|
||||
- case. case => d k [H ck]. left. exists (d,k). split ; last by []. rewrite ltn_neqAle in H.
|
||||
by apply (proj2 (andP H)).
|
||||
- move => X. case: (isEps_dec (c n)).
|
||||
+ case => d E. left. exists (d,n). by rewrite E.
|
||||
+ move => E. right => k. case_eq (k == n) => EE ; first by rewrite (eqP EE).
|
||||
move => L. rewrite ltnS in L. rewrite leq_eqVlt in L. rewrite EE in L. by apply X.
|
||||
Defined.
|
||||
|
||||
Definition createchain (c:natO =-> lift_ordType) (n:nat)
|
||||
(X:{dk: (D*nat)%type | match dk with (d,k) => k<n /\ c k =-= Val d end}) :
|
||||
(natO =-> D) :=
|
||||
match X with | exist (d,k) (conj Hk Hck) => @makechainm c k d Hck
|
||||
end.
|
||||
|
||||
Lemma makechainworks : forall (c:natO =-> lift_ordType) k dk (H2:c k =-= Val dk) i (d:D), c (k+i)%N =-= (Val d) -> makechain H2 i =-= d.
|
||||
move => c k dk Hck i d e.
|
||||
apply vinj.
|
||||
rewrite <- e.
|
||||
apply Oeq_sym.
|
||||
unfold makechain.
|
||||
rewrite <- extractworks.
|
||||
simpl. by rewrite addnC.
|
||||
Qed.
|
||||
|
||||
Lemma eqDLeq : forall d d', d =-= d' -> Val d =-= (Val d').
|
||||
move => d d'. by case ; split ; apply DLle_leVal.
|
||||
Qed.
|
||||
|
||||
Lemma predeq : forall x, x =-= pred x.
|
||||
move => x. case: x ; last by [].
|
||||
move => x. simpl. by rewrite <-( @pred_nth_eq 1 (Eps x)).
|
||||
Qed.
|
||||
|
||||
(* Just realized I should have had Val as a morphism ages ago *)
|
||||
Add Parametric Morphism : (@Val D)
|
||||
with signature (@tset_eq D) ==> (@tset_eq lift_ordType)
|
||||
as Val_eq_compat.
|
||||
intros.
|
||||
apply eqDLeq.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
|
||||
Add Parametric Morphism : (@Val D)
|
||||
with signature (@Ole D) ++> (@Ole lift_ordType)
|
||||
as Val_le_compat.
|
||||
intros.
|
||||
apply DLle_leVal.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma DLle_Val_exists_eq : forall y d, Val d <= y -> exists d', y =-= Val d' /\ d <= d'.
|
||||
intros y d H; inversion H.
|
||||
exists d'.
|
||||
split.
|
||||
symmetry.
|
||||
apply Val_exists_pred_eq.
|
||||
exists n; trivial.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
Lemma DLvalgetchain: forall (c:natO =-> lift_ordType) k d, c k =-= Val d -> exists c':natO =-> D, forall n, c(k+n)%N =-= Val (c' n).
|
||||
intros c k d chk.
|
||||
exists (makechainm chk).
|
||||
intros n.
|
||||
destruct (allvalsfromhere n chk) as [d' [chd ld]].
|
||||
refine (Oeq_trans chd _).
|
||||
apply eqDLeq.
|
||||
apply (Oeq_sym (makechainworks chk chd)).
|
||||
Qed.
|
||||
|
||||
Hint Immediate DLle_leVal.
|
||||
|
||||
Lemma eta_mono : monotonic (@Val D).
|
||||
move => x y L. by rewrite -> L.
|
||||
Qed.
|
||||
|
||||
Definition eta_m : D =-> lift_ordType := Eval hnf in mk_fmono eta_mono.
|
||||
|
||||
End Lift_ord.
|
||||
|
||||
Implicit Arguments eta_m [D].
|
||||
|
||||
Section Lift_cpo.
|
||||
|
||||
Variable D:cpoType.
|
||||
|
||||
CoFixpoint DL_lubn (c:natO =-> lift_ordType D) (n:nat) : lift_ordType D :=
|
||||
match fValIdx c n with inleft X => Val (lub (createchain X))
|
||||
| inright _ => Eps (DL_lubn (cpred c) (S n))
|
||||
end.
|
||||
|
||||
Lemma DL_lubn_inv : forall (c:natO =-> lift_ordType D) (n:nat), DL_lubn c n =
|
||||
match fValIdx c n with inleft X => Val (lub (createchain X))
|
||||
| inright _ => Eps (DL_lubn (cpred c) (S n))
|
||||
end.
|
||||
intros; rewrite (DL_inv (DL_lubn c n)).
|
||||
simpl; case (fValIdx c n); trivial.
|
||||
Qed.
|
||||
|
||||
Lemma chainluble : forall (c:natO =-> lift_ordType D) k dk (Hkdk : c k =-= Val dk) k' dk' (Hkdk': c k' =-= Val dk'),
|
||||
lub (makechainm Hkdk) <= lub (makechainm Hkdk').
|
||||
intros;apply lub_le.
|
||||
intro n.
|
||||
assert (Hkn := allvalsfromhere n Hkdk).
|
||||
destruct Hkn as (dkn,(Hckn,Hdkdkn)).
|
||||
setoid_replace (makechainm Hkdk n) with (dkn).
|
||||
destruct (allvalsfromhere (n+k) Hkdk') as [dkk'n [Hckk'n Hddd]].
|
||||
apply Ole_trans with dkk'n.
|
||||
apply vleinj; rewrite <- Hckn; rewrite <- Hckk'n.
|
||||
apply fmonotonic. rewrite addnC. by apply: (leq_addl k').
|
||||
setoid_replace dkk'n with (makechainm Hkdk' (n+k)%N).
|
||||
apply le_lub.
|
||||
unfold makechainm.
|
||||
apply Oeq_sym; apply makechainworks.
|
||||
assumption.
|
||||
unfold makechainm; apply makechainworks.
|
||||
assumption.
|
||||
Qed.
|
||||
|
||||
Lemma chainlubinv : forall (c:natO =-> lift_ordType D) k dk (Hkdk : c k =-= Val dk) k' dk' (Hkdk': c k' =-= Val dk'),
|
||||
lub (makechainm Hkdk) =-= lub (makechainm Hkdk').
|
||||
intros;split;apply chainluble.
|
||||
Qed.
|
||||
|
||||
Lemma pred_lubn_Val : forall (d:D)(n k p:nat) (c:natO =-> lift_ordType D),
|
||||
(n <k+p)%nat -> pred_nth (c n) k = Val d
|
||||
-> exists d', DL_lubn c p =-= Val d'.
|
||||
move => d n. elim.
|
||||
- move => p c L E. rewrite (DL_lubn_inv c p). simpl in E.
|
||||
case: (fValIdx c p).
|
||||
+ case.
|
||||
* move => [d' k] [H1 H2]. simpl. by exists (lub (makechainm H2)).
|
||||
* move => X. specialize (X _ L). by rewrite E in X.
|
||||
+ move => k IH p c L E. rewrite (DL_lubn_inv c p).
|
||||
case: (fValIdx c p).
|
||||
* unfold createchain. case. case. move => d' k'. case => LL EE.
|
||||
by exists (lub (makechainm EE)).
|
||||
* move => X. rewrite addSn in L. rewrite <- addnS in L.
|
||||
have A: (pred_nth (cpred c n) k = Val d) by simpl ; rewrite <- pred_nth_Sn_acc.
|
||||
specialize (IH (S p) (cpred c) L A).
|
||||
case: IH => d' IH. exists d'. rewrite <- IH. by rewrite <- eq_Eps.
|
||||
Qed.
|
||||
|
||||
Lemma cpredsamelub : forall (c:natO =-> lift_ordType D) k dk (H1:c k =-= Val dk) (H2: (cpred c) k =-= Val dk),
|
||||
lub (makechainm H1) =-= lub (makechainm H2).
|
||||
intros; split.
|
||||
apply: lub_le; intro.
|
||||
setoid_replace (makechainm H1 n) with (makechainm H2 n).
|
||||
apply le_lub.
|
||||
assert (Hkn := allvalsfromhere n H1).
|
||||
destruct Hkn as (dkn,(Hckn,Hdkdkn)).
|
||||
unfold makechainm; simpl.
|
||||
setoid_rewrite (makechainworks H1 Hckn).
|
||||
assert (HH : ((cpred c) (k+n)%N =-= Val dkn)).
|
||||
unfold cpred; simpl.
|
||||
rewrite <- Hckn.
|
||||
apply Oeq_sym; apply predeq.
|
||||
setoid_rewrite (makechainworks H2 HH).
|
||||
auto.
|
||||
|
||||
apply: lub_le; intro.
|
||||
setoid_replace (makechainm H2 n) with (makechainm H1 n).
|
||||
apply le_lub.
|
||||
destruct (allvalsfromhere n H2) as (dkn,(Hckn,Hdkdkn)).
|
||||
unfold makechainm; simpl.
|
||||
setoid_rewrite (makechainworks H2 Hckn).
|
||||
assert (HH : ((c) (k+n)%N =-= Val dkn)).
|
||||
unfold cpred in Hckn; simpl in Hckn.
|
||||
rewrite <- Hckn; apply predeq.
|
||||
setoid_rewrite (makechainworks H1 HH); auto.
|
||||
Qed.
|
||||
|
||||
Lemma attempt2 : forall kl (c:natO =-> lift_ordType D) p d, pred_nth (DL_lubn c p) kl = Val d
|
||||
-> exists k, exists dk, exists Hkdk:c k =-= Val dk, d =-= lub (makechainm Hkdk).
|
||||
elim.
|
||||
- simpl. move => c p d E. rewrite -> DL_lubn_inv in E.
|
||||
case: (fValIdx c p) E ; last by [].
|
||||
move => s. case => e. rewrite <- e. clear d e.
|
||||
case: s. case => d n [L E]. simpl. exists n. exists d. by exists E.
|
||||
- move => k IH c p d E. simpl in E. unfold createchain in E.
|
||||
have HH:=IH (cpred c) (S p) d.
|
||||
case: (fValIdx c p) E.
|
||||
+ case. case => d' n [P Q] E.
|
||||
specialize (IH c p d). rewrite E in IH. apply IH. by apply pred_nth_val.
|
||||
+ move => X E. case: (HH E) => n [x [P Q]].
|
||||
exists n. exists x. have PP:=P. simpl in PP. have P0:=Oeq_trans (predeq _) P. exists P0. rewrite -> Q.
|
||||
by apply (Oeq_sym (cpredsamelub P0 P)).
|
||||
Qed.
|
||||
|
||||
(* first just wrap pred_lubn_Val to use equality *)
|
||||
Lemma chainVallubnVal : forall (c:natO =-> lift_ordType D) n d p, c n =-= Val d -> exists d', DL_lubn c p =-= Val d'.
|
||||
intros.
|
||||
destruct (eqValpred H) as (k,(d'',(Hcnk,Hdd''))).
|
||||
apply (pred_lubn_Val (n:=n) (k:=k+n+1) (p:=p) (c:=c) (d:=d'')).
|
||||
apply ltn_addr. rewrite <- addnA. apply ltn_addl. rewrite addnS. rewrite addn0. by apply ltnSn.
|
||||
rewrite addnS. rewrite addn0. rewrite <- addnS.
|
||||
rewrite pred_nth_sum.
|
||||
rewrite Hcnk. by [].
|
||||
Qed.
|
||||
|
||||
Lemma chainVallubnlub : forall (c:natO =-> lift_ordType D) n d (H : c n =-= Val d) p, exists d', DL_lubn c p =-= Val d' /\ d' =-= lub (makechainm H).
|
||||
intros.
|
||||
destruct (chainVallubnVal p H) as (d',HH).
|
||||
exists d'.
|
||||
split. assumption.
|
||||
destruct (eqValpred HH) as (k,(d'',(Hk,Hd'd''))).
|
||||
destruct (attempt2 Hk) as (p',(dp',(Hp',Heq))).
|
||||
rewrite Hd'd''.
|
||||
rewrite -> Heq.
|
||||
apply chainlubinv.
|
||||
Qed.
|
||||
|
||||
|
||||
Definition DL_lub (c:natO =-> lift_ordType D) := DL_lubn c 1.
|
||||
|
||||
|
||||
Lemma DL_lub_upper : forall c:natO =-> lift_ordType D, forall n, c n <= DL_lub c.
|
||||
intros c n. simpl. apply: DLless_cond.
|
||||
intros d C.
|
||||
destruct (chainVallubnlub C 1) as [d' [Ln L]].
|
||||
unfold DL_lub. rewrite -> Ln.
|
||||
assert (X:=makechainworks C). specialize (X 0 d). rewrite addn0 in X.
|
||||
specialize (X C). apply Ole_trans with (y:=Val d). by auto.
|
||||
assert (pred_nth (Val d') 0 = Val d') as Y by auto.
|
||||
apply (DLleVal Y). rewrite -> L.
|
||||
apply Ole_trans with (y:=(makechainm (c:=c) (k:=n) (d:=d) C) 0). simpl. by auto.
|
||||
by apply le_lub.
|
||||
Qed.
|
||||
|
||||
Lemma DL_lub_least : forall (c : natO =-> lift_ordType D) a,
|
||||
(forall n, c n <= a) -> DL_lub c <= a.
|
||||
intros c a C. simpl.
|
||||
apply: DLless_cond. intros d X.
|
||||
destruct (eqValpred X) as [n [dd [P Q]]].
|
||||
destruct (attempt2 P) as [m [d' [cm Y]]].
|
||||
apply Ole_trans with (y:= (Val dd)).
|
||||
refine (proj2 (Val_exists_pred_eq _)). exists n. auto.
|
||||
assert (Z:= C m). rewrite -> cm in Z.
|
||||
destruct (DLvalval Z) as [mm [e [R S]]].
|
||||
apply (DLleVal R). rewrite -> Y.
|
||||
apply lub_le.
|
||||
intros nn.
|
||||
assert (A:=makechainworks cm).
|
||||
specialize (A nn).
|
||||
clear S Y Z.
|
||||
destruct (allvalsfromhere nn cm) as [nnv [S T]].
|
||||
specialize (A _ S). rewrite -> A. clear A.
|
||||
assert (a =-= Val e) as E.
|
||||
apply Oeq_sym.
|
||||
apply (Val_exists_pred_eq). exists mm. apply R.
|
||||
assert (Z := C (m+nn)%N).
|
||||
rewrite -> E in Z. rewrite -> S in Z. apply (vleinj Z).
|
||||
Qed.
|
||||
|
||||
|
||||
(** *** Declaration of the lifted cpo *)
|
||||
|
||||
Lemma DLCpoAxiom : CPO.axiom DL_lub.
|
||||
move => x y n.
|
||||
split ; first by apply DL_lub_upper.
|
||||
by apply DL_lub_least.
|
||||
Qed.
|
||||
|
||||
Canonical Structure lift_cpoMixin := CpoMixin DLCpoAxiom.
|
||||
Canonical Structure lift_cpoType := Eval hnf in CpoType lift_cpoMixin.
|
||||
|
||||
Lemma lubval: forall (c:natO =-> lift_cpoType) d, lub c =-= Val d ->
|
||||
exists k, exists d', c k =-= Val d' /\ d' <= d.
|
||||
intros c d l.
|
||||
destruct l as [lubval1 lubval2].
|
||||
destruct (DLle_Val_exists_pred lubval2) as [k xx].
|
||||
destruct xx as [f' [lubval lf]].
|
||||
assert (lub c =-= Val f') as lubval'.
|
||||
apply: Ole_antisym.
|
||||
eapply Ole_trans. apply lubval1. by apply: DLle_leVal.
|
||||
simpl.
|
||||
apply: (DLleVal). apply lubval. apply Ole_refl.
|
||||
assert (f' <= d).
|
||||
destruct lubval' as [v1 v2].
|
||||
assert (Val f' <= Val d).
|
||||
apply (Ole_trans v2 lubval1).
|
||||
apply (DLleVal_leq H).
|
||||
assert (d =-= f') as feq. auto.
|
||||
clear H lubval1 lubval2 lf.
|
||||
|
||||
assert (xx:=attempt2 lubval).
|
||||
destruct xx as [n [f0 [chnval lf']]].
|
||||
exists n. exists f0. split. apply chnval.
|
||||
assert (w:=makechainworks chnval).
|
||||
specialize (w 0 f0). rewrite addn0 in w.
|
||||
specialize (w chnval).
|
||||
destruct w as [_ L2].
|
||||
refine (Ole_trans L2 (Ole_trans (le_lub (makechainm chnval) 0) _)).
|
||||
rewrite <- lf'. by auto.
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Eta
|
||||
==========================================================================*)
|
||||
|
||||
(** printing eta %\ensuremath{\eta}% *)
|
||||
|
||||
Lemma eta_cont : continuous (@eta_m D).
|
||||
move => c. simpl. unfold lub. simpl. unfold DL_lub. rewrite -> DL_lubn_inv. simpl.
|
||||
apply: DLle_leVal.
|
||||
apply lub_le => n. apply: (Ole_trans _ (le_lub _ n)). simpl.
|
||||
rewrite -> makechainworks. by apply Ole_refl. by apply Oeq_refl.
|
||||
Qed.
|
||||
|
||||
Definition eta : D =-> lift_cpoType := Eval hnf in mk_fcont eta_cont.
|
||||
|
||||
End Lift_cpo.
|
||||
|
||||
Implicit Arguments eta [D].
|
||||
|
||||
(** printing _BOT %\LIFTED% *)
|
||||
Notation "x '_BOT'" := (lift_cpoType x) (at level 28).
|
||||
|
||||
Lemma liftOrdPointed (D:ordType) : Pointed.axiom (DL_bot D).
|
||||
move => x. by apply DL_bot_least.
|
||||
Qed.
|
||||
|
||||
Canonical Structure liftOrdPointedMixin D := PointedMixin (@liftOrdPointed D).
|
||||
Canonical Structure liftOrdPointedType D := Eval hnf in PointedType (liftOrdPointedMixin D).
|
||||
|
||||
Lemma liftCpoPointed (D:cpoType) : Pointed.axiom (DL_bot D).
|
||||
move => x. by apply DL_bot_least.
|
||||
Qed.
|
||||
|
||||
Canonical Structure liftCppoMixin D := PointedMixin (@liftCpoPointed D).
|
||||
Canonical Structure liftCppoType D := Eval hnf in CppoType (liftCppoMixin D).
|
||||
Canonical Structure liftCpoPointedType (D:cpoType) := Eval hnf in @PointedType (D _BOT) (liftCppoMixin D).
|
||||
|
||||
Lemma liftFunPointed C (D:pointedType) : Pointed.axiom (@fmon_cte C D PBot).
|
||||
move => f x. simpl. by apply leastP.
|
||||
Qed.
|
||||
|
||||
Canonical Structure funOrdPointedMixin C D := PointedMixin (@liftFunPointed C D).
|
||||
Canonical Structure funOrdPointedType C D := Eval hnf in PointedType (funOrdPointedMixin C D).
|
||||
|
||||
Lemma PBot_app D (E:cppoType) : forall d, (PBot:D -=> E) d =-= PBot.
|
||||
move => d. by simpl.
|
||||
Qed.
|
||||
|
||||
Lemma PBot_incon (D:cpoType) (x:D) : eta x <= PBot -> False.
|
||||
move => incon.
|
||||
inversion incon. subst. clear x H1 incon.
|
||||
elim: n H0.
|
||||
- simpl. unfold Pointed.least. simpl. by rewrite -> DL_bot_eq.
|
||||
- move => n IH. unfold Pointed.least. simpl. by apply IH.
|
||||
Qed.
|
||||
|
||||
Lemma PBot_incon_eq (D:cpoType) (x:D) : eta x =-= PBot -> False.
|
||||
intros [incon _].
|
||||
apply (PBot_incon incon).
|
||||
Qed.
|
885
papers/domains-2012/PredomRec.v
Executable file
885
papers/domains-2012/PredomRec.v
Executable file
|
@ -0,0 +1,885 @@
|
|||
(**********************************************************************************
|
||||
* PredomRec.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
Require Export Categories PredomCore NSetoid.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
(** printing T0 %\ensuremath{T_0}% *)
|
||||
(** printing T1 %\ensuremath{T_1}% *)
|
||||
(** printing T2 %\ensuremath{T_2}% *)
|
||||
(** printing T3 %\ensuremath{T_3}% *)
|
||||
(** printing T4 %\ensuremath{T_4}% *)
|
||||
(** printing T5 %\ensuremath{T_5}% *)
|
||||
(** printing T6 %\ensuremath{T_6}% *)
|
||||
|
||||
Open Scope C_scope.
|
||||
|
||||
(*=Enriched *)
|
||||
Module BaseCat.
|
||||
Record mixin_of (O:Type) (M:O -> O -> cppoType) := Mixin
|
||||
{ catm :> Category.mixin_of M;
|
||||
terminal :> CatTerminal.mixin_of (Category.Pack catm O);
|
||||
comp : forall X Y Z : O, M Y Z * M X Y =-> M X Z;
|
||||
comp_comp : forall X Y Z m m', comp X Y Z (m,m') =-= Category.tcomp catm m m'
|
||||
}. (*CLEAR*)
|
||||
Notation class_of := mixin_of (only parsing).
|
||||
(*CLEARED*)
|
||||
Section ClassDef.
|
||||
Definition base2 T (M:T -> T -> cppoType) (c:class_of M) := CatTerminal.Class c. (*CLEAR*)
|
||||
Local Coercion base2 : class_of >-> CatTerminal.class_of.
|
||||
|
||||
Structure cat := Pack {object; morph : object -> object -> cppoType ; _ : class_of morph; _ : Type}.
|
||||
Local Coercion object : cat >-> Sortclass.
|
||||
Local Coercion morph : cat >-> Funclass.
|
||||
Definition class cT := let: Pack _ _ c _ := cT return class_of (@morph cT) in c.
|
||||
Definition unpack (K:forall O (M:O -> O -> cppoType) (c:class_of M), Type)
|
||||
(k : forall O (M: O -> O -> cppoType) (c : class_of M), K _ _ c) (cT:cat) :=
|
||||
let: Pack _ M c _ := cT return @K _ _ (class cT) in k _ _ c.
|
||||
Definition repack cT : _ -> Type -> cat := let k T M c p := p c in unpack k cT.
|
||||
Definition pack T M c := @Pack T M c T.
|
||||
(*CLEARED*)
|
||||
Definition terminalCat (cT:cat) : terminalCat := CatTerminal.Pack (class cT) cT.
|
||||
Definition catType (cT:cat) : catType := Category.Pack (class cT) cT. (*CLEAR*)
|
||||
Definition ordType (cT:cat) (X Y:cT) : ordType := PreOrd.Pack (CPPO.class (morph X Y)) (cT X Y).
|
||||
Definition cpoType (cT:cat) (X Y:cT) : cpoType := CPO.Pack (CPPO.class (morph X Y)) (cT X Y).
|
||||
(*CLEARED*)
|
||||
Definition cppoType (cT:cat) (X Y:cT) : cppoType :=
|
||||
CPPO.Pack (CPPO.class (morph X Y)) (cT X Y). (*CLEAR*)
|
||||
Definition pointedType (cT:cat) (X Y:cT) : pointedType := Pointed.Pack (CPPO.class (morph X Y)) (cT X Y).
|
||||
End ClassDef.
|
||||
Module Import Exports.
|
||||
Coercion terminalCat : cat >-> CatTerminal.cat.
|
||||
Coercion base2 : class_of >-> CatTerminal.class_of.
|
||||
Coercion object : cat >-> Sortclass.
|
||||
Coercion morph : cat >-> Funclass.
|
||||
Canonical Structure BaseCat.terminalCat.
|
||||
Canonical Structure BaseCat.catType.
|
||||
Canonical Structure BaseCat.ordType.
|
||||
Canonical Structure BaseCat.cpoType.
|
||||
Canonical Structure BaseCat.pointedType.
|
||||
Canonical Structure BaseCat.cppoType.
|
||||
Notation cppoMorph := BaseCat.cppoType. (*CLEAR*)
|
||||
Notation cppoECat := BaseCat.cat.
|
||||
Notation CppoECatMixin := BaseCat.Mixin.
|
||||
(*CLEARED*)
|
||||
Notation CppoECatType := BaseCat.pack.
|
||||
(*CLEARED*)
|
||||
End Exports.
|
||||
End BaseCat. (*CLEAR*)
|
||||
Export BaseCat.Exports.
|
||||
(*CLEARED*)
|
||||
(*=End *)
|
||||
|
||||
Definition getmorph (C:cppoECat) (X Y : C) (f:CPPO.cpoType (@BaseCat.cppoType C X Y)) : C X Y := f.
|
||||
Coercion getmorph : cppoECat >-> Funclass.
|
||||
|
||||
Definition ccomp (cT:cppoECat) (X Y Z : cT) : cppoMorph Y Z * cppoMorph X Y =-> cppoMorph X Z := BaseCat.comp (BaseCat.class cT) X Y Z.
|
||||
|
||||
Lemma ccomp_eq (cT:cppoECat) (X Y Z : cT) (f : cT Y Z) (f' : cT X Y) : ccomp X Y Z (f,f') =-= f << f'.
|
||||
unfold Category.comp. unfold ccomp. simpl. case:cT X Y Z f f'. simpl.
|
||||
move => O M. case. simpl.
|
||||
move => C T comp ce T' X Y Z f g. by apply: ce.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (M:cppoECat) (X Y Z:M) : (@Category.comp M X Y Z)
|
||||
with signature (@Ole (cppoMorph Y Z)) ++> (@Ole (cppoMorph X Y)) ++> (@Ole (cppoMorph X Z))
|
||||
as comp_le_compat.
|
||||
move => f f' l g g' l'.
|
||||
have e:= (ccomp_eq f g). have e':=ccomp_eq f' g'. rewrite <- e. rewrite <- e'. apply: fmonotonic.
|
||||
by split.
|
||||
Qed.
|
||||
|
||||
Section Msolution.
|
||||
Variable M:cppoECat.
|
||||
|
||||
Open Scope C_scope.
|
||||
Open Scope S_scope.
|
||||
|
||||
(*=BiFunctor *)
|
||||
Record BiFunctor : Type := mk_functor
|
||||
{ ob : M -> M -> M ;
|
||||
morph : forall (T0 T1 T2 T3 : M), (cppoMorph T1 T0) * (cppoMorph T2 T3) =->
|
||||
(cppoMorph (ob T0 T2) (ob T1 T3)) ;
|
||||
morph_comp: forall T0 T1 T2 T3 T4 T5 (f:M T4 T1) (g:M T3 T5) (h:M T1 T0)
|
||||
(k:M T2 T3), getmorph (morph T1 T4 T3 T5 (f,g)) << (morph T0 T1 T2 T3 (h, k)) =-=
|
||||
morph _ _ _ _ (h << f, g << k) ;
|
||||
morph_id : forall T0 T1, morph T0 T0 T1 T1 (Id:M _ _, Id:M _ _) =-= (Id : M _ _)}.
|
||||
(*=End *)
|
||||
|
||||
Definition eppair T0 T1 (f:M T0 T1) (g:M T1 T0) :=
|
||||
g << f =-= Id /\ f << g <= @Category.id M _.
|
||||
|
||||
(*=Tower *)
|
||||
Record Tower : Type := mk_tower
|
||||
{ tobjects : nat -> M;
|
||||
tmorphisms : forall i, (tobjects (S i)) =-> (tobjects i);
|
||||
tmorphismsI : forall i, (tobjects i) =-> (tobjects (S i));
|
||||
teppair : forall i, eppair (tmorphismsI i) (tmorphisms i)}.
|
||||
(*=End *)
|
||||
|
||||
(*=Cone *)
|
||||
Record Cone (To:Tower) : Type := mk_basecone
|
||||
{ tcone :> M;
|
||||
mcone : forall i, tcone =-> (tobjects To i);
|
||||
mconeCom : forall i, tmorphisms To i << mcone (S i) =-= mcone i }.
|
||||
(*=End *)
|
||||
|
||||
Record CoCone (To:Tower) : Type := mk_basecocone
|
||||
{ tcocone :> M;
|
||||
mcocone : forall i, (tobjects To i) =-> tcocone;
|
||||
mcoconeCom : forall i, mcocone (S i) << tmorphismsI To i =-= mcocone i }.
|
||||
|
||||
(*=CoLimit *)
|
||||
Record CoLimit (To:Tower) : Type := mk_basecolimit
|
||||
{ lcocone :> CoCone To;
|
||||
colimitExists : forall (A:CoCone To), (tcocone lcocone) =-> (tcocone A) ;
|
||||
colimitCom : forall (A:CoCone To), forall n, mcocone A n =-= colimitExists A << mcocone lcocone n;
|
||||
colimitUnique : forall (A:CoCone To) (h: (tcocone lcocone) =-> (tcocone A))
|
||||
(C:forall n, mcocone A n =-= h << mcocone lcocone n), h =-= colimitExists A }.
|
||||
(*=End *)
|
||||
|
||||
(*=Limit *)
|
||||
Record Limit (To:Tower) : Type := mk_baselimit
|
||||
{lcone :> Cone To;
|
||||
limitExists : forall (A:Cone To), (tcone A) =-> (tcone lcone);
|
||||
limitCom : forall (A:Cone To), forall n, mcone A n =-= mcone lcone n << limitExists A;
|
||||
limitUnique : forall (A:Cone To) (h: (tcone A) =-> (tcone lcone))
|
||||
(C:forall n, mcone A n =-= mcone lcone n << h), h =-= limitExists A }.
|
||||
|
||||
Variable L:forall T:Tower, Limit T.
|
||||
(*=End *)
|
||||
|
||||
Section RecursiveDomains.
|
||||
|
||||
Variable F : BiFunctor.
|
||||
|
||||
Add Parametric Morphism C X Y Z W : (fun (x: Y =-> X) (y: Z =-> W) => @morph C X Y Z W (pair x y))
|
||||
with signature (@tset_eq _) ==> (@tset_eq _) ==> (@tset_eq _)
|
||||
as morph_eq_compat.
|
||||
move => x x' e y y' e'. apply: fmon_stable. by split ; split ;case: e ; case: e'.
|
||||
Qed.
|
||||
|
||||
Lemma BiFuncRtoR: forall (T0 T1:M) (f: T0 =-> T1) (g: T1 =-> T0), eppair f g ->
|
||||
eppair (morph F _ _ _ _ (g,f)) (morph F _ _ _ _ (f,g)).
|
||||
move => A B f g. case => e r. split.
|
||||
- rewrite -> (morph_comp F f g g f). rewrite -> (morph_eq_compat F e e).
|
||||
by rewrite morph_id.
|
||||
- rewrite morph_comp. rewrite -> (pair_le_compat r r). by rewrite morph_id.
|
||||
Qed.
|
||||
|
||||
Section Tower.
|
||||
|
||||
Variable DTower : Tower.
|
||||
|
||||
Lemma lt_gt_dec n m : (n < m) + (m = n) + (m < n).
|
||||
case: (ltngtP m n) => e.
|
||||
- by right.
|
||||
- left. by left.
|
||||
- left. by right.
|
||||
Qed.
|
||||
|
||||
Lemma leqI m n : (m <= n)%N = false -> n <= m.
|
||||
move => I. have t:=(leq_total m n). rewrite I in t. by apply t.
|
||||
Qed.
|
||||
|
||||
Lemma lt_subK n m : m < n -> n = (n - m + m)%N.
|
||||
move => l. apply sym_eq. apply subnK. have x:= (leqnSn m). by apply (leq_trans x l).
|
||||
Qed.
|
||||
|
||||
(*=Projections *)
|
||||
Definition Diter_coercion n m : n = m ->
|
||||
M (tobjects DTower n) (tobjects DTower m) :=
|
||||
fun e => eq_rect_r (fun n : nat => M (tobjects DTower n) _) Id e.
|
||||
|
||||
Fixpoint Projection_nm m (n:nat) :
|
||||
(tobjects DTower (n+m)) =-> (tobjects DTower m) :=
|
||||
match n with
|
||||
| O => Id
|
||||
| S n => Projection_nm m n << tmorphisms DTower (n+m)%N
|
||||
end.
|
||||
|
||||
Fixpoint Injection_nm m (n:nat) :
|
||||
(tobjects DTower m) =-> (tobjects DTower (n+m)) :=
|
||||
match n with
|
||||
| O => Id
|
||||
| S n => tmorphismsI DTower (n+m) << Injection_nm m n
|
||||
end.
|
||||
|
||||
Definition t_nm n m : (tobjects DTower n) =-> (tobjects DTower m) :=
|
||||
match (lt_gt_dec m n) with
|
||||
| inl (inl ee) =>
|
||||
(Projection_nm m (n - m)%N) << (Diter_coercion (lt_subK ee))
|
||||
| inl (inr ee) => Diter_coercion ee
|
||||
| inr ee => Diter_coercion (sym_eq (lt_subK ee)) << Injection_nm n (m - n)
|
||||
end.
|
||||
(*=End *)
|
||||
|
||||
Lemma Diter_coercion_simpl: forall n (Eq:n = n), Diter_coercion Eq = Id.
|
||||
intros n Eq. by rewrite (eq_irrelevance Eq (refl_equal _)).
|
||||
Qed.
|
||||
|
||||
Lemma Proj_left_comp : forall k m, Projection_nm m (S k) =-=
|
||||
tmorphisms DTower m << Projection_nm (S m) k << Diter_coercion (sym_eq (addnS k m)).
|
||||
elim.
|
||||
- move => m. simpl.
|
||||
rewrite comp_idL. rewrite comp_idR.
|
||||
rewrite (Diter_coercion_simpl). by rewrite -> comp_idR.
|
||||
- move => n IH m. simpl. simpl in IH. rewrite -> IH.
|
||||
have P:=(sym_eq (addnS n m)). rewrite (eq_irrelevance (sym_eq (addnS n m)) P).
|
||||
have Q:=(sym_eq (addnS n.+1 m)). rewrite (eq_irrelevance (sym_eq (addnS n.+1 m)) Q).
|
||||
move: P Q. unfold addn. simpl. fold addn. move => P. generalize P. rewrite P. clear P.
|
||||
move => P Q. do 2 rewrite Diter_coercion_simpl.
|
||||
do 2 rewrite comp_idR. clear P Q. by rewrite comp_assoc.
|
||||
Qed.
|
||||
|
||||
Lemma Inject_right_comp: forall k m, Injection_nm m (S k) =-=
|
||||
Diter_coercion (addnS k m) << Injection_nm (S m) k << tmorphismsI DTower m.
|
||||
elim.
|
||||
- move => m. simpl. do 2 rewrite comp_idR. rewrite (Diter_coercion_simpl). by rewrite -> comp_idL.
|
||||
- move => n IH m. simpl. simpl in IH. rewrite -> IH. clear IH.
|
||||
have P:= (addnS n.+1 m). have Q:=(addnS n m). rewrite -> (eq_irrelevance _ P).
|
||||
rewrite -> (eq_irrelevance _ Q). move: P Q. unfold addn. simpl. fold addn.
|
||||
move => P Q. generalize P Q. rewrite <- Q. clear P Q. move => P Q. do 2 rewrite Diter_coercion_simpl.
|
||||
do 2 rewrite comp_idL. by rewrite comp_assoc.
|
||||
Qed.
|
||||
|
||||
Lemma Diter_coercion_comp: forall x y z (Eq1:x = y) (Eq2 : y = z),
|
||||
Diter_coercion Eq2 << Diter_coercion Eq1 =-= Diter_coercion (trans_equal Eq1 Eq2).
|
||||
intros x y z Eq1. generalize Eq1. rewrite Eq1. clear x Eq1.
|
||||
intros Eq1 Eq2. generalize Eq2. rewrite <- Eq2. clear Eq2. intros Eq2.
|
||||
repeat (rewrite Diter_coercion_simpl). by rewrite -> comp_idL.
|
||||
Qed.
|
||||
|
||||
Lemma Diter_coercion_eq n m (e e': n = m) : Diter_coercion e =-= Diter_coercion e'.
|
||||
by rewrite (eq_irrelevance e e').
|
||||
Qed.
|
||||
|
||||
Lemma lt_gt_dec_lt n m : n < m -> exists e, lt_gt_dec n m = inl _ (inl _ e).
|
||||
move => e. case_eq (lt_gt_dec n m) ; first case.
|
||||
- move => i _. by exists i.
|
||||
- move => f. rewrite f in e. by rewrite ltnn in e.
|
||||
- move => f. have a:=leq_trans (leqnSn _) f. have Ff:=leq_ltn_trans a e. by rewrite ltnn in Ff.
|
||||
Qed.
|
||||
|
||||
Lemma lt_gt_dec_gt n m : m < n -> exists e, lt_gt_dec n m = (inr _ e).
|
||||
move => e. case_eq (lt_gt_dec n m) ; first case.
|
||||
- move => i _. have a:=leq_trans (leqnSn _) i. have Ff:=leq_ltn_trans a e. by rewrite ltnn in Ff.
|
||||
- move => f. rewrite f in e. by rewrite ltnn in e.
|
||||
- move => i _. by exists i.
|
||||
Qed.
|
||||
|
||||
Lemma lt_gt_dec_eq n m : forall e : m = n, lt_gt_dec n m = (inl _ (inr _ e)).
|
||||
move => e ; generalize e ; rewrite e. clear e m => e.
|
||||
case_eq (lt_gt_dec n n) ; first case.
|
||||
- move => i _. have f:=i. by rewrite ltnn in f.
|
||||
- move => f _. by rewrite (eq_irrelevance e f).
|
||||
- move => i _. have f:=i. by rewrite ltnn in f.
|
||||
Qed.
|
||||
|
||||
Lemma t_nmProjection: forall n m, t_nm n m =-= tmorphisms DTower m << t_nm n (S m).
|
||||
move => n m. unfold t_nm. case (lt_gt_dec m n) ; first case.
|
||||
- case (lt_gt_dec m.+1 n) ; first case.
|
||||
+ move => l l'. rewrite comp_assoc.
|
||||
have X:=Proj_left_comp (n - m.+1) m. have XX:= addnS (n - m.+1) m.
|
||||
have Y:=comp_eq_compat X (tset_refl (Diter_coercion XX)).
|
||||
repeat (rewrite <- comp_assoc in Y). rewrite -> Diter_coercion_comp in Y.
|
||||
rewrite Diter_coercion_simpl in Y. rewrite -> comp_idR in Y. rewrite <- Y.
|
||||
rewrite <- comp_assoc. rewrite Diter_coercion_comp.
|
||||
have P:=(trans_equal (lt_subK l) XX). have Q:=(lt_subK l').
|
||||
rewrite -> (eq_irrelevance _ P). rewrite -> (eq_irrelevance _ Q).
|
||||
generalize P. have e:= (ltn_subS l'). have ee:(n - m).-1 = (n - m.+1). by rewrite e.
|
||||
rewrite <- ee. clear ee e P. have e:0 < n - m. rewrite <- ltn_add_sub. by rewrite addn0.
|
||||
move => P. have PP:n = ((n - m).-1.+1 + m)%N. rewrite (ltn_predK e). apply Q.
|
||||
rewrite (eq_irrelevance P PP). move: PP. rewrite (ltn_predK e). clear. move => PP.
|
||||
by rewrite (eq_irrelevance Q PP).
|
||||
+ move => e. generalize e. rewrite e. clear n e. move => e i. have x:= subSnn m.
|
||||
have a:=(lt_subK (n:=m.+1) (m:=m) i). rewrite (eq_irrelevance (lt_subK (n:=m.+1) (m:=m) i) a).
|
||||
move: a. rewrite x. clear x. move => a. rewrite (eq_irrelevance e a). simpl. by rewrite comp_idL.
|
||||
+ move => i j. have Ff:=leq_trans i j. by rewrite ltnn in Ff.
|
||||
- move => e. generalize e. rewrite e. clear n e. move => e. destruct (lt_gt_dec_gt (leqnn (S m))) as [p A].
|
||||
rewrite A. clear A.
|
||||
have ee:=subSnn m. have x:=(sym_eq (lt_subK (n:=m.+1) (m:=m) p)).
|
||||
rewrite <- (eq_irrelevance x (sym_eq (lt_subK (n:=m.+1) (m:=m) p))). move: x. rewrite ee. simpl.
|
||||
move => x. do 2 rewrite (Diter_coercion_simpl). rewrite comp_idL. rewrite comp_idR.
|
||||
unfold addn. simpl. by rewrite -> (proj1 (teppair DTower m)).
|
||||
- move => i. have l:=leq_trans i (leqnSn m). destruct (lt_gt_dec_gt l) as [p A]. rewrite A. clear A l.
|
||||
have x:=(sym_eq (lt_subK (n:=m.+1) (m:=n) p)). rewrite (eq_irrelevance (sym_eq (lt_subK (n:=m.+1) (m:=n) p)) x).
|
||||
move: x. rewrite leq_subS ; last by []. simpl. move => x. do 2 rewrite comp_assoc.
|
||||
refine (comp_eq_compat _ (tset_refl (Injection_nm n (m-n)))). have y:=(sym_eq (lt_subK (n:=m) (m:=n) i)).
|
||||
rewrite (eq_irrelevance (sym_eq (lt_subK (n:=m) (m:=n) i)) y). have a:=(trans_eq (sym_eq (addSn _ _)) x).
|
||||
rewrite (eq_irrelevance x a).
|
||||
apply tset_trans with (y:=(tmorphisms DTower m << Diter_coercion a) << tmorphismsI DTower (m - n + n)) ; last by [].
|
||||
move: a. generalize y. rewrite -> y. clear. move => y a. do 2 rewrite Diter_coercion_simpl.
|
||||
rewrite comp_idR. by rewrite (proj1 (teppair DTower m)).
|
||||
Qed.
|
||||
|
||||
Lemma t_nn_ID: forall n, t_nm n n =-= Id.
|
||||
intros n.
|
||||
unfold t_nm. rewrite (lt_gt_dec_eq (refl_equal n)). by rewrite Diter_coercion_simpl.
|
||||
Qed.
|
||||
|
||||
Lemma t_nmProjection2 n m : (m <= n) % nat -> t_nm (S n) m =-= (t_nm n m) << tmorphisms DTower n.
|
||||
move => nm.
|
||||
assert (exists x, n - m = x) as X by (exists (n - m) ; auto).
|
||||
destruct X as [x X]. elim: x n m X nm.
|
||||
- move => n m E EE. have e:=eqn_leq n m. rewrite EE in e. have l:(n <= m)%N. unfold leq. by rewrite E.
|
||||
rewrite l in e. have ee:=eqP e. rewrite ee. rewrite t_nn_ID. rewrite comp_idL.
|
||||
rewrite -> t_nmProjection. rewrite t_nn_ID. by rewrite comp_idR.
|
||||
- move => nm IH. case ; first by []. move => n m e l. have ll:n.+1 - m > 0 by rewrite e. rewrite subn_gt0 in ll.
|
||||
rewrite leq_subS in e ; last by []. specialize (IH n m). have ee: n - m = nm by auto.
|
||||
rewrite -> (IH ee ll).
|
||||
clear IH. have l0:=leq_trans ll (leqnSn _). destruct (lt_gt_dec_lt l0) as [p A].
|
||||
unfold t_nm. rewrite A. clear A. case_eq (lt_gt_dec m n) ; first case.
|
||||
+ move => p' _. have x:=(lt_subK (n:=n.+2) (m:=m) p). rewrite (eq_irrelevance (lt_subK (n:=n.+2) (m:=m) p) x).
|
||||
move: x. have x:n.+2 - m = (n-m).+2. rewrite leq_subS ; last by []. rewrite leq_subS ; by [].
|
||||
rewrite x. clear x. move => x. simpl. do 4 rewrite <- comp_assoc.
|
||||
refine (comp_eq_compat (tset_refl (Projection_nm m _)) _).
|
||||
have x':=(lt_subK (n:=n) (m:=m) p'). rewrite (eq_irrelevance (lt_subK (n:=n) (m:=m) p') x').
|
||||
have x1:((n - m).+2 + m)%N = ((n - m) + m).+2 by [].
|
||||
have x2:=trans_eq x x1.
|
||||
apply tset_trans with (y:=tmorphisms DTower (n - m + m) << (tmorphisms DTower (S(n - m + m)) << Diter_coercion x2)) ;
|
||||
first by simpl ; rewrite (eq_irrelevance x2 x).
|
||||
move: x2. generalize x' ; rewrite <- x'. clear x' x x1 p p'. move => x x'. do 2 rewrite (Diter_coercion_simpl).
|
||||
rewrite comp_idR. by rewrite -> comp_idL.
|
||||
+ move => e'. generalize e'. move: p. rewrite e'. clear e' n l ll e l0 ee. move => p e _.
|
||||
rewrite (Diter_coercion_simpl). rewrite comp_idL.
|
||||
have x:=(lt_subK (n:=m.+2) (m:=m) p).
|
||||
have ee:(m.+2 - m) = (m - m).+2. rewrite leq_subS ; last by []. rewrite leq_subS ; by [].
|
||||
rewrite (eq_irrelevance (lt_subK (n:=m.+2) (m:=m) p) x). generalize x. rewrite ee. clear ee. rewrite subnn.
|
||||
move => xx. clear x e p. simpl. rewrite comp_idL.
|
||||
rewrite (Diter_coercion_simpl). by rewrite -> comp_idR.
|
||||
+ move => f. have lx:=leq_trans f ll. by rewrite ltnn in lx.
|
||||
Qed.
|
||||
|
||||
Lemma t_nmEmbedding: forall n i, t_nm n i =-= (t_nm (S n) i) << tmorphismsI DTower n.
|
||||
intros n m. unfold t_nm. case (lt_gt_dec m n) ; first case.
|
||||
- move => i. have l:=leq_trans i (leqnSn _). destruct (lt_gt_dec_lt l) as [p A]. rewrite A. clear A l.
|
||||
have x:=((lt_subK p)). rewrite (eq_irrelevance ((lt_subK p)) x).
|
||||
move: x. rewrite leq_subS ; last by []. simpl. move => x. do 2 rewrite <- comp_assoc.
|
||||
refine (comp_eq_compat (tset_refl (Projection_nm _ _)) _). have y:=((lt_subK i)).
|
||||
rewrite (eq_irrelevance ((lt_subK i)) y). have a:=(trans_eq x ((addSn _ _))).
|
||||
rewrite (eq_irrelevance x a).
|
||||
apply tset_trans with (y:=tmorphisms DTower (n - m + m) << (Diter_coercion a << tmorphismsI DTower n)) ; last by [].
|
||||
move: a. generalize y. rewrite <- y. clear. move => y a. do 2 rewrite Diter_coercion_simpl.
|
||||
rewrite comp_idL. by rewrite <- (proj1 (teppair DTower n)).
|
||||
- move => e. generalize e. rewrite e. clear n e. move => e. destruct (lt_gt_dec_lt (leqnn (S m))) as [p A].
|
||||
rewrite A. clear A.
|
||||
have ee:=subSnn m. have x:=((lt_subK (n:=m.+1) (m:=m) p)).
|
||||
rewrite <- (eq_irrelevance x ((lt_subK (n:=m.+1) (m:=m) p))). move: x. rewrite ee. simpl.
|
||||
move => x. do 2 rewrite (Diter_coercion_simpl). rewrite comp_idR. rewrite comp_idL.
|
||||
by rewrite -> (proj1 (teppair DTower m)).
|
||||
- case (lt_gt_dec m n.+1) ; first case.
|
||||
+ move => i j. have Ff:=leq_trans i j. by rewrite ltnn in Ff.
|
||||
+ move => e. generalize e. rewrite <- e. clear m e. move => e i. have x:= subSnn n.
|
||||
have a:=(lt_subK i). rewrite (eq_irrelevance (lt_subK i) a).
|
||||
move: a. rewrite x. clear x. move => a. rewrite (eq_irrelevance (sym_eq a) e).
|
||||
simpl. by rewrite comp_idR.
|
||||
+ move => l l'. rewrite <- comp_assoc. have X:=Inject_right_comp (m - n.+1) n. have XX:= sym_eq (addnS (m - n.+1) n).
|
||||
have Y:=comp_eq_compat (tset_refl (Diter_coercion XX)) X. rewrite -> comp_assoc in Y.
|
||||
rewrite -> (comp_assoc (Injection_nm _ _)) in Y. rewrite -> Diter_coercion_comp in Y.
|
||||
rewrite Diter_coercion_simpl in Y. rewrite -> comp_idL in Y. rewrite <- Y.
|
||||
rewrite comp_assoc. rewrite Diter_coercion_comp.
|
||||
have ee:=(trans_equal XX (sym_eq (lt_subK l))). rewrite -> (eq_irrelevance _ ee).
|
||||
clear XX X Y. have e:m - n.+1 = (m - n).-1. by rewrite (ltn_subS l').
|
||||
move: ee. rewrite -> e. have ll:0 < (m - n). rewrite <- ltn_add_sub. by rewrite addn0.
|
||||
move => Q.
|
||||
have PP:((m - n).-1.+1 + n)%N = m. rewrite -> (ltn_predK ll). by rewrite (sym_eq (lt_subK l')).
|
||||
rewrite (eq_irrelevance Q PP). move: PP. have P:=(sym_eq (lt_subK l')). rewrite -> (eq_irrelevance _ P).
|
||||
move: P. clear Q. clear e. move => P Q. move: P. rewrite -{1 2 3 4} (ltn_predK ll).
|
||||
move => P. by rewrite (eq_irrelevance P Q).
|
||||
Qed.
|
||||
|
||||
Lemma t_nmEmbedding2: forall n m, (n <= m) % nat -> t_nm n (S m) =-= tmorphismsI DTower m << t_nm n m.
|
||||
move => n m nm.
|
||||
assert (exists x, m - n = x) as X by (eexists ; apply refl_equal).
|
||||
destruct X as [x X]. elim: x n m X nm.
|
||||
- move => n m E EE. have e:=eqn_leq n m. rewrite EE in e. have l:(m <= n)%N. unfold leq. by rewrite E.
|
||||
rewrite l in e. have ee:=eqP e. rewrite ee. rewrite -> t_nn_ID. rewrite comp_idR.
|
||||
rewrite -> t_nmEmbedding. rewrite t_nn_ID. by rewrite comp_idL.
|
||||
- move => nm IH. move => n. case ; first by []. move => m e l. have ll:m.+1 - n > 0 by rewrite e. rewrite subn_gt0 in ll.
|
||||
rewrite leq_subS in e ; last by []. specialize (IH n m). have ee:m - n = nm by auto.
|
||||
rewrite (IH ee ll). clear IH. have l0:=leq_trans ll (leqnSn _). destruct (lt_gt_dec_gt l0) as [p A].
|
||||
unfold t_nm. rewrite A. clear A. case_eq (lt_gt_dec m n) ; first case.
|
||||
+ move => f. have lx:=leq_trans f ll. by rewrite ltnn in lx.
|
||||
+ move => e'. generalize e'. move: p. rewrite e'. clear e' n l ll e l0 ee. move => p e _.
|
||||
rewrite (Diter_coercion_simpl). rewrite comp_idR. have x:=(lt_subK p).
|
||||
have ee:(m.+2 - m) = (m - m).+2. rewrite leq_subS ; last by []. rewrite leq_subS ; by [].
|
||||
rewrite (eq_irrelevance (lt_subK (n:=m.+2) (m:=m) p) x). generalize x. rewrite ee. clear ee. rewrite subnn.
|
||||
move => xx. clear x e p. simpl. rewrite (Diter_coercion_simpl). rewrite -> comp_idL. by rewrite comp_idR.
|
||||
+ move => p' _. have x:=(lt_subK p). rewrite (eq_irrelevance (lt_subK p) x).
|
||||
move: x. have x:m.+2 - n = (m-n).+2. rewrite leq_subS ; last by []. rewrite leq_subS ; by [].
|
||||
rewrite x. clear x. move => x. simpl. do 4 rewrite -> comp_assoc. refine (comp_eq_compat _ (tset_refl (Injection_nm _ _))).
|
||||
have x':=sym_eq(lt_subK p'). rewrite (eq_irrelevance (sym_eq (lt_subK _)) x').
|
||||
have x1:((m - n).+2 + n)%N = ((m - n) + n).+2 by [].
|
||||
have x2:=sym_eq (trans_eq x x1).
|
||||
apply tset_trans with (y:=(Diter_coercion x2 << tmorphismsI DTower (S (m - n + n))) << tmorphismsI DTower ((m - n + n))) ;
|
||||
first by simpl ; rewrite (eq_irrelevance x2 (sym_eq x)).
|
||||
move: x2. generalize x' ; rewrite x'. clear x' x x1 p p'. move => x x'. do 2 rewrite (Diter_coercion_simpl).
|
||||
rewrite -> comp_idR. by rewrite comp_idL.
|
||||
Qed.
|
||||
|
||||
Lemma t_nm_EP i n : (i <= n)%N -> eppair (t_nm i n) (t_nm n i).
|
||||
have e:exists x, x == n - i by eexists.
|
||||
case: e => x e. elim: x i n e.
|
||||
- move => i n e l. have ee:i = n. apply anti_leq. rewrite l. simpl. unfold leq. by rewrite <- (eqP e).
|
||||
rewrite -> ee. split.
|
||||
+ rewrite t_nn_ID. by rewrite comp_idL.
|
||||
+ apply Oeq_le. rewrite t_nn_ID. by rewrite comp_idL.
|
||||
- move => x IH i. case ; first by rewrite sub0n.
|
||||
move => n e l. have ll:(i <= n)%N. have aa:0 < n.+1 - i. by rewrite <- (eqP e). rewrite subn_gt0 in aa. by apply aa.
|
||||
clear l. specialize (IH i n). rewrite (leq_subS ll) in e. specialize (IH e ll). split.
|
||||
+ rewrite (t_nmProjection2 ll). rewrite -> (t_nmEmbedding2 ll). rewrite comp_assoc.
|
||||
rewrite <- (comp_assoc (tmorphismsI _ _)). rewrite (proj1 (teppair DTower n)). rewrite comp_idR. by apply (proj1 IH).
|
||||
+ rewrite -> (t_nmEmbedding2 ll). rewrite (t_nmProjection2 ll).
|
||||
rewrite comp_assoc. rewrite <- (comp_assoc (t_nm n i)).
|
||||
rewrite -> (comp_le_compat (comp_le_compat (Ole_refl _) (proj2 IH)) (Ole_refl _)).
|
||||
rewrite comp_idR. by rewrite -> (proj2 (teppair DTower n)).
|
||||
Qed.
|
||||
|
||||
Definition coneN (n:nat) : Cone DTower.
|
||||
exists (tobjects DTower n) (t_nm n).
|
||||
move => m. simpl. by rewrite ->(t_nmProjection n m).
|
||||
Defined.
|
||||
|
||||
Lemma coneCom_l (X:Cone DTower) n i : (i <= n)%nat -> mcone X i =-= t_nm n i << mcone X n.
|
||||
move => l.
|
||||
assert (exists x, n - i = x) as E by (eexists ; auto).
|
||||
destruct E as [x E]. elim: x n i l E.
|
||||
- move => n i l E. have A:=@anti_leq n i. rewrite l in A. unfold leq in A. rewrite E in A.
|
||||
rewrite andbT in A. specialize (A (eqtype.eq_refl _)). rewrite A. clear l E A n.
|
||||
rewrite t_nn_ID. by rewrite comp_idL.
|
||||
- move => x IH n i l E. rewrite -> (comp_eq_compat (t_nmProjection n i) (tset_refl _)).
|
||||
rewrite <- comp_assoc. specialize (IH n i.+1). have l0:i < n by rewrite <- subn_gt0 ; rewrite E.
|
||||
have ee:n - i.+1 = x. have Y:= subn_gt0 i n. rewrite E in Y. rewrite ltn_subS in E ; first by auto. by rewrite <- Y.
|
||||
specialize (IH l0 ee). rewrite <- IH.
|
||||
by rewrite -> (mconeCom X i).
|
||||
Qed.
|
||||
|
||||
Lemma coconeCom_l (X:CoCone DTower) n i : (n <= i)%nat -> mcocone X n =-= mcocone X i << (t_nm n i).
|
||||
move => l.
|
||||
assert (exists x, i - n = x) as E by (eexists ; auto).
|
||||
destruct E as [x E]. elim: x n i l E.
|
||||
- move => n i l E. have A:=@anti_leq n i. rewrite l in A. unfold leq in A. rewrite E in A.
|
||||
specialize (A is_true_true). rewrite A. rewrite -> t_nn_ID. by rewrite comp_idR.
|
||||
- move => x IH n i l E. rewrite -> t_nmEmbedding.
|
||||
rewrite -> comp_assoc. specialize (IH n.+1 i).
|
||||
have ll:n < i by rewrite <- subn_gt0 ; rewrite E.
|
||||
have ee:i - n.+1 = x by have Y:= subn_gt0 n i ; rewrite E in Y ;
|
||||
rewrite ltn_subS in E ; try rewrite <- Y ; by auto.
|
||||
rewrite <- (IH ll ee). by apply (tset_sym (mcoconeCom X n)).
|
||||
Qed.
|
||||
|
||||
Lemma chainPEm (C:CoCone DTower) (C':Cone DTower) : monotonic (fun n => (mcocone C n << mcone C' n) : cppoMorph _ _).
|
||||
move => i j l.
|
||||
have E:exists x, j - i = x by exists (j - i).
|
||||
case: E l => x. elim: x i j.
|
||||
- move => i j e l. have l':(j <= i)%N. unfold leq. by rewrite e.
|
||||
have ee: j <= i <= j. rewrite l'. by apply l.
|
||||
by rewrite (anti_leq ee).
|
||||
- move => x IH i. case ; first by rewrite sub0n.
|
||||
move => j l _. specialize (IH i.+1 j.+1). rewrite subSS in IH. rewrite <- IH.
|
||||
+ have EE:= comp_eq_compat (mcoconeCom C i) (mconeCom C' i).
|
||||
rewrite <- comp_assoc in EE. rewrite -> (comp_assoc (mcone C' _)) in EE.
|
||||
have ll:= (proj2 (teppair DTower i)).
|
||||
rewrite -> (proj2 EE).
|
||||
apply: comp_le_compat ; first by [].
|
||||
rewrite -> (comp_le_compat ll (Ole_refl _)). by rewrite comp_idL.
|
||||
+ clear IH. have ll:(i < j.+1)%N. rewrite <- subn_gt0. by rewrite l.
|
||||
rewrite (@leq_subS i j ll) in l. by auto.
|
||||
+ have ll:(i < j.+1)%N. rewrite <- subn_gt0. by rewrite l. by apply ll.
|
||||
Qed.
|
||||
|
||||
Definition chainPE (C:CoCone DTower) (C':Cone DTower) : natO =-> cppoMorph _ _ := mk_fmono (@chainPEm C C').
|
||||
|
||||
End Tower.
|
||||
|
||||
(*=Diter *)
|
||||
Fixpoint Diter (n:nat) :=
|
||||
match n return M with | O => One | S n => ob F (Diter n) (Diter n) end.
|
||||
Fixpoint Injection (n:nat) : (Diter n) =-> (Diter (S n)) :=
|
||||
match n with
|
||||
| O => PBot
|
||||
| S n => morph F _ _ _ _ (Projection n, Injection n)
|
||||
end with Projection (n:nat) : (Diter (S n)) =-> (Diter n) :=
|
||||
match n with
|
||||
| O => terminal_morph _
|
||||
| S n => morph F _ _ _ _ (Injection n,Projection n)
|
||||
end.
|
||||
Variable comp_left_strict : forall (X Y Z:M) (f:M X Y),
|
||||
(PBot:Y =-> Z) << f =-= (PBot:X =-> Z).
|
||||
Lemma eppair_IP: forall n, eppair (Injection n) (Projection n).
|
||||
(*=End *)
|
||||
elim.
|
||||
- split ; first by apply:terminal_unique.
|
||||
simpl. rewrite -> (comp_left_strict _ (terminal_morph _)). by apply: leastP.
|
||||
- move => n IH. split.
|
||||
+ simpl. rewrite -> (morph_comp F (Injection n) (Projection n) (Projection n) (Injection n)).
|
||||
rewrite -> (morph_eq_compat F (proj1 IH) (proj1 IH)). by rewrite morph_id.
|
||||
+ simpl. rewrite morph_comp.
|
||||
rewrite -> (pair_le_compat (proj2 IH) (proj2 IH)). by rewrite morph_id.
|
||||
Qed.
|
||||
|
||||
(*=DTower *)
|
||||
Definition DTower := mk_tower eppair_IP.
|
||||
(*=End *)
|
||||
|
||||
(*=DInf *)
|
||||
Definition DInf : M := tcone (L DTower).
|
||||
(*=End *)
|
||||
|
||||
Definition Embeddings n : (tobjects DTower n) =-> DInf := limitExists (L DTower) (coneN DTower n).
|
||||
|
||||
Definition Projections n : DInf =-> Diter n := mcone (L DTower) n.
|
||||
|
||||
Lemma coCom i : Embeddings i.+1 << Injection i =-= Embeddings i.
|
||||
unfold Embeddings.
|
||||
rewrite -> (limitUnique (h:=(Embeddings (S i) << Injection _ : (tcone (coneN DTower i)) =-> DInf))) ; first by [].
|
||||
move => m. simpl. unfold Embeddings. rewrite -> comp_assoc. have e:= (limitCom (L DTower) (coneN DTower i.+1) m).
|
||||
simpl in e. rewrite <- e. by rewrite -> (t_nmEmbedding DTower).
|
||||
Qed.
|
||||
|
||||
Definition DCoCone : CoCone DTower := @mk_basecocone _ DInf Embeddings coCom.
|
||||
|
||||
Lemma retract_EP n : Projections n << Embeddings n =-= Id.
|
||||
unfold eppair.
|
||||
unfold Projections. unfold Embeddings.
|
||||
- rewrite <- (limitCom (L DTower) (coneN DTower n) n). simpl. by rewrite -> t_nn_ID.
|
||||
Qed.
|
||||
|
||||
Lemma emp: forall i j, Projections i << Embeddings j =-= t_nm DTower j i.
|
||||
intros i j. have A:= leq_total i j. case_eq (i <= j)%N ; move => X ; rewrite X in A.
|
||||
- unfold Projections. rewrite -> (comp_eq_compat (coneCom_l (L DTower) X) (tset_refl _)).
|
||||
rewrite <- comp_assoc. rewrite -> (comp_eq_compat (tset_refl _) (retract_EP j)). by rewrite -> comp_idR.
|
||||
- simpl in A. have e:= (coconeCom_l DCoCone A). simpl in e. rewrite -> e.
|
||||
rewrite -> comp_assoc. rewrite -> (retract_EP i). by rewrite comp_idL.
|
||||
Qed.
|
||||
|
||||
Definition chainPEc (C:CoCone DTower) : natO =-> (cppoMorph DInf (tcocone C)) := chainPE C (L DTower).
|
||||
|
||||
Lemma lub_comp_left (X Y:M) (f:X =-> Y) Z (c:natO =-> cppoMorph Y Z) :
|
||||
(lub c:M _ _) << f =-= lub ( (exp_fun (ccomp _ _ _ << <|pi2,pi1|>) f:ordCatType _ _) << c).
|
||||
rewrite <- (ccomp_eq (lub c) f). by rewrite <- (lub_comp_eq _ c).
|
||||
Qed.
|
||||
|
||||
Lemma lub_comp_right (Y Z:M) (f:Y =-> Z) X (c:natO =-> cppoMorph X Y) :
|
||||
f << lub c =-= lub ( (exp_fun (ccomp _ _ _) f:ordCatType _ _) << c).
|
||||
rewrite <- (ccomp_eq f (lub c)). by rewrite <- lub_comp_eq.
|
||||
Qed.
|
||||
|
||||
Lemma lub_comp_both (X Y Z:M) (c:natO =-> cppoMorph X Y) (c':natO =-> cppoMorph Y Z) :
|
||||
(lub c':M _ _) << lub c =-= lub ((ccomp _ _ _ : ordCatType _ _) << <|c', c|>).
|
||||
rewrite <- lub_comp_eq.
|
||||
rewrite {3} /lub. simpl. unfold prod_lub. rewrite prod_fun_fst. rewrite prod_fun_snd.
|
||||
by rewrite -> ccomp_eq.
|
||||
Qed.
|
||||
|
||||
Lemma EP_id : (lub (chainPEc DCoCone) : M _ _) =-= Id.
|
||||
rewrite <- (tset_sym (@limitUnique _ (L DTower) (L DTower) Id (fun n => tset_sym (comp_idR _)))).
|
||||
apply: limitUnique. move => n. rewrite -> lub_comp_right.
|
||||
rewrite -> (lub_lift_left _ n). apply: Ole_antisym. rewrite <- (le_lub _ O). simpl. rewrite addn0.
|
||||
have e:= comp_eq_compat (retract_EP n) (tset_refl (mcone (L DTower) n)).
|
||||
rewrite -> comp_idL in e. unfold Projections in e. rewrite <- comp_assoc in e.
|
||||
rewrite ccomp_eq. by rewrite e.
|
||||
|
||||
apply: lub_le. move => i. simpl. rewrite ccomp_eq. rewrite -> comp_assoc.
|
||||
rewrite -> (emp n (n+i)). by rewrite (coneCom_l _ (leq_addr i n)).
|
||||
Qed.
|
||||
|
||||
Lemma chainPE_simpl X n : chainPEc X n = mcocone X n << Projections n.
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition DCoLimit : CoLimit DTower.
|
||||
exists DCoCone (fun C => lub (chainPEc C)).
|
||||
- move => C n. simpl. rewrite -> lub_comp_left. apply Ole_antisym.
|
||||
+ apply: (Ole_trans _ (le_lub _ n)).
|
||||
simpl. rewrite -> ccomp_eq.
|
||||
have e:= comp_eq_compat (tset_refl (mcocone C n)) (emp n n). rewrite -> t_nn_ID in e.
|
||||
rewrite -> comp_idR in e. rewrite -> comp_assoc in e. by rewrite -> e.
|
||||
+ rewrite -> (lub_lift_left _ n). apply: lub_le => k.
|
||||
simpl. rewrite -> ccomp_eq.
|
||||
have e:=comp_eq_compat (tset_refl (mcocone C (n+k))) (emp (n+k) n). rewrite -> comp_assoc in e.
|
||||
rewrite <- (coconeCom_l C (leq_addr k n)) in e. by rewrite e.
|
||||
- move => C h X. rewrite <- (comp_idR h). rewrite <- (EP_id).
|
||||
apply tset_trans with (y:=exp_fun (ccomp _ _ _) h (lub (chainPEc DCoCone))) ; first by apply: (tset_sym (ccomp_eq _ _)).
|
||||
rewrite -> (@lub_comp_eq _ _ ( (exp_fun (ccomp _ _ _) h)) (chainPEc DCoCone)).
|
||||
apply: lub_eq_compat. apply fmon_eq_intro => n. simpl. specialize (X n).
|
||||
rewrite ccomp_eq.
|
||||
rewrite comp_assoc. by rewrite <- X.
|
||||
Defined.
|
||||
|
||||
Definition ECoCone : CoCone DTower.
|
||||
exists (ob F DInf DInf) (fun i => (morph F _ _ _ _ (Projections i, Embeddings i):M _ _) << Injection i).
|
||||
move => i. simpl. rewrite -> (morph_comp F (Projections i.+1) (Embeddings i.+1) (Projection i) (Injection i)).
|
||||
by rewrite -> (morph_eq_compat F (mconeCom (L DTower) i) (mcoconeCom DCoCone i)).
|
||||
Defined.
|
||||
|
||||
Definition FCone : Cone DTower.
|
||||
exists (ob F DInf DInf) (fun n => Projection n << morph F _ _ _ _ (Embeddings n, Projections n)).
|
||||
move => i. refine (comp_eq_compat (tset_refl _) _). unfold Projections.
|
||||
rewrite <- (morph_eq_compat F (coCom i) (mconeCom (L DTower) i)). simpl.
|
||||
apply: morph_comp.
|
||||
Defined.
|
||||
|
||||
Definition chainFPE (C:CoCone DTower) := chainPE C FCone.
|
||||
|
||||
Lemma subS_leq j i m : j.+1 - i == m.+1 -> i <= j.
|
||||
move => X. have a:=subn_gt0 i (j.+1). have x:=ltn0Sn m.
|
||||
rewrite <- (eqP X) in x. rewrite a in x. by [].
|
||||
Qed.
|
||||
|
||||
Lemma morph_tnm n m : morph F _ _ _ _ (t_nm DTower m n, t_nm DTower n m) =-= t_nm DTower n.+1 m.+1.
|
||||
have t:= (leq_total n m). case_eq (n <= m)%N ; move => l ; rewrite l in t.
|
||||
clear t. have M':exists x, m - n == x by eexists ; by []. destruct M' as [x M'].
|
||||
elim: x m n l M'.
|
||||
- move => j i l e. rewrite subn_eq0 in e.
|
||||
have A:= @anti_leq i j. rewrite l in A. rewrite e in A. specialize (A is_true_true).
|
||||
rewrite A. rewrite -> (morph_eq_compat F (t_nn_ID DTower j) (t_nn_ID DTower j)).
|
||||
rewrite -> morph_id. by rewrite -> t_nn_ID.
|
||||
- move => m IH. case.
|
||||
+ move=> i l e. rewrite leqn0 in l. have ee:=eqP l. rewrite ee.
|
||||
rewrite -> (morph_eq_compat F (t_nn_ID DTower O) (t_nn_ID DTower O)).
|
||||
rewrite -> morph_id. by rewrite -> t_nn_ID.
|
||||
+ move => j i e l. specialize (IH j i). have l':=l. have l0:=subS_leq l'. rewrite (leq_subS l0) in l.
|
||||
specialize (IH l0 l). clear l l' e.
|
||||
have e:= morph_eq_compat F (t_nmProjection2 DTower l0) (t_nmEmbedding2 DTower l0).
|
||||
rewrite <- morph_comp in e. rewrite -> e. rewrite -> IH. clear e.
|
||||
have a:= (tset_sym (t_nmEmbedding2 DTower _)).
|
||||
specialize (a i.+1 j.+1 l0). by apply a.
|
||||
simpl in t. clear l.
|
||||
have M':exists x, n - m == x by eexists ; by []. destruct M' as [x M'].
|
||||
elim: x m n t M'.
|
||||
- move => j i l e. rewrite subn_eq0 in e.
|
||||
have A:= @anti_leq i j. rewrite l in A. rewrite e in A. specialize (A is_true_true).
|
||||
rewrite A. rewrite -> (morph_eq_compat F (t_nn_ID DTower j) (t_nn_ID DTower j)).
|
||||
rewrite -> morph_id. by rewrite -> t_nn_ID.
|
||||
- move => m IH i. case.
|
||||
+ move=> l e. rewrite leqn0 in l. have ee:=eqP l. rewrite ee.
|
||||
rewrite -> (morph_eq_compat F (t_nn_ID DTower O) (t_nn_ID DTower O)).
|
||||
rewrite -> morph_id. by rewrite -> t_nn_ID.
|
||||
+ move => j e l. specialize (IH i j). have l':=l. have l0:=subS_leq l'. rewrite (leq_subS l0) in l.
|
||||
specialize (IH l0 l). clear l l' e.
|
||||
have e:= morph_eq_compat F (t_nmEmbedding2 DTower l0) (t_nmProjection2 DTower l0).
|
||||
rewrite <-morph_comp in e. rewrite -> e. simpl. rewrite -> IH.
|
||||
clear e. have a:= (tset_sym (t_nmProjection2 DTower _)).
|
||||
specialize (a j.+1 i.+1 l0). by apply a.
|
||||
Qed.
|
||||
|
||||
Lemma ECoLCom : forall (A : CoCone DTower) (n : nat),
|
||||
mcocone A n =-= (lub (chainFPE A):M _ _) << mcocone ECoCone n.
|
||||
move => C n. simpl.
|
||||
rewrite -> lub_comp_left. apply Ole_antisym.
|
||||
* rewrite <- (le_lub _ n). simpl. rewrite -> ccomp_eq. do 2 rewrite comp_assoc.
|
||||
rewrite <- (comp_assoc (morph F _ _ _ _ _ : M _ _)).
|
||||
rewrite -> (morph_comp F (Embeddings n) (Projections n) (Projections n) (Embeddings n)).
|
||||
rewrite -> (morph_eq_compat F (tset_trans (emp n n) (t_nn_ID _ _))
|
||||
(tset_trans (emp n n) (t_nn_ID _ _))).
|
||||
rewrite -> morph_id. rewrite comp_idR. rewrite <- comp_assoc. rewrite -> (proj1 (teppair DTower n)).
|
||||
by rewrite comp_idR.
|
||||
* rewrite -> (lub_lift_left _ n). apply: lub_le => k. simpl. apply Oeq_le.
|
||||
rewrite -> ccomp_eq. do 2 rewrite comp_assoc. rewrite <- (comp_assoc (morph F _ _ _ _ _:M _ _)).
|
||||
rewrite -> (morph_comp F (Embeddings (n+k)) (Projections (n+k)) (Projections n) (Embeddings n)).
|
||||
rewrite -> (morph_eq_compat F (emp n (n+k)) (emp (n+k) n)).
|
||||
rewrite morph_tnm. rewrite <- comp_assoc. rewrite <- (t_nmEmbedding DTower n (n+k).+1).
|
||||
rewrite <- comp_assoc. rewrite <- (t_nmProjection DTower n (n+k)).
|
||||
by rewrite <- (coconeCom_l C (leq_addr k n)).
|
||||
Qed.
|
||||
|
||||
Open Scope D_scope.
|
||||
Open Scope C_scope.
|
||||
|
||||
Lemma pair_lub (D E:cpoType) (c:natO =-> D) (c':natO =-> E) : @tset_eq (CPO.setoidType _) (lub c,lub c') (lub (<|c,c'|>)).
|
||||
rewrite {3} /lub. simpl. unfold prod_lub.
|
||||
split ; split ; simpl ; apply lub_le_compat ; apply Oeq_le ; try (by rewrite prod_fun_fst) ; by rewrite prod_fun_snd.
|
||||
Qed.
|
||||
|
||||
Lemma CoLimitUnique : forall (A : CoCone DTower) h,
|
||||
(forall n : nat, mcocone A n =-= h << mcocone ECoCone n) ->
|
||||
h =-= lub (chainFPE A).
|
||||
move => C h X. rewrite <- (comp_idR h).
|
||||
have a:= morph_eq_compat F EP_id EP_id. rewrite -> morph_id in a.
|
||||
rewrite <- a. clear a.
|
||||
have b:= (pair_lub (chainPEc DCoCone) (chainPEc DCoCone)).
|
||||
have S:= (fmon_stable (morph F _ _ _ _) b). rewrite -> (Oeq_trans S (lub_comp_eq _ _)).
|
||||
clear b S. rewrite -> lub_comp_right. rewrite -> (lub_lift_left (chainFPE C) 1).
|
||||
apply: lub_eq_compat. apply fmon_eq_intro => n. simpl.
|
||||
rewrite -> (ccomp_eq h (morph F _ _ _ _ (Embeddings n << mcone (L DTower) n,
|
||||
Embeddings n << mcone (L DTower) n))).
|
||||
specialize (X n.+1). simpl in X.
|
||||
rewrite -> X. rewrite (morph_comp F (Projections n.+1) (Embeddings n.+1) (Projection n) (Injection n)).
|
||||
rewrite -> (morph_comp F (Injection n) (Projection n) (Embeddings (1 + n)) (Projections (1 + n))).
|
||||
rewrite -> (morph_eq_compat F (mconeCom (L DTower) n) (mcoconeCom DCoCone n)).
|
||||
rewrite -> (morph_eq_compat F (mcoconeCom DCoCone n) (mconeCom (L DTower) n)).
|
||||
rewrite <- comp_assoc.
|
||||
by rewrite -> (morph_comp F (mcone (L DTower) n) (mcocone DCoCone n) (mcocone DCoCone n) (mcone (L DTower) n)).
|
||||
Qed.
|
||||
|
||||
Definition ECoLimit : CoLimit DTower.
|
||||
exists ECoCone (fun C => lub (chainFPE C)).
|
||||
by apply ECoLCom.
|
||||
by apply CoLimitUnique.
|
||||
Defined.
|
||||
|
||||
(*=Iso *)
|
||||
Definition Fold : (ob F DInf DInf) =-> DInf := (* ... *) (*CLEAR*) colimitExists ECoLimit DCoCone. (*CLEARED*)
|
||||
Definition Unfold : DInf =-> (ob F DInf DInf) := (*CLEAR*) colimitExists DCoLimit ECoCone. (*CLEARED*)
|
||||
|
||||
Lemma FU_id : Fold << Unfold =-= Id. Proof.
|
||||
apply tset_trans with (y:=colimitExists DCoLimit DCoLimit).
|
||||
- refine (@colimitUnique _ DCoLimit DCoLimit _ _). move => i. unfold Fold. unfold Unfold.
|
||||
simpl. rewrite -> lub_comp_both. rewrite -> lub_comp_left. apply Ole_antisym.
|
||||
* apply: (Ole_trans _ (le_lub _ i)).
|
||||
simpl. apply Oeq_le. rewrite ccomp_eq. rewrite ccomp_eq.
|
||||
do 3 rewrite comp_assoc. rewrite <- (comp_assoc (morph F _ _ _ _ _ :M _ _)).
|
||||
rewrite -> (morph_comp F (Embeddings i) (Projections i) (Projections i) (Embeddings i)).
|
||||
rewrite -> (morph_eq_compat F (tset_trans (emp i i) (t_nn_ID _ _)) (tset_trans (emp i i) (t_nn_ID _ _))).
|
||||
rewrite -> morph_id. rewrite comp_idR. rewrite <- (comp_assoc (Injection i)).
|
||||
rewrite -> (proj1 (teppair DTower i)). rewrite comp_idR. rewrite <- comp_assoc.
|
||||
rewrite -> (emp i i). rewrite t_nn_ID. by rewrite comp_idR.
|
||||
* rewrite -> (lub_lift_left _ i). apply: lub_le => k. simpl.
|
||||
apply Oeq_le. rewrite -> ccomp_eq. rewrite ccomp_eq.
|
||||
do 3 rewrite comp_assoc. rewrite <- (comp_assoc (morph F _ _ _ _ _:M _ _)).
|
||||
rewrite -> (morph_comp F (Embeddings (i + k)) (Projections (i + k)) (Projections (i + k)) (Embeddings (i + k))).
|
||||
rewrite (morph_eq_compat F (tset_trans (emp _ _) (t_nn_ID _ _)) (tset_trans (emp _ _) (t_nn_ID _ _))).
|
||||
rewrite -> morph_id. rewrite comp_idR. rewrite <- (comp_assoc (Injection (i+k))).
|
||||
rewrite -> (proj1 (teppair DTower (i+k))). rewrite comp_idR. rewrite <- comp_assoc.
|
||||
rewrite -> (emp (i+k) (i)).
|
||||
by apply (tset_sym (@coconeCom_l DTower DCoCone _ _ (leq_addr k i))).
|
||||
- apply tset_sym. apply: (@colimitUnique _ DCoLimit DCoLimit _ _). move => n. simpl. by rewrite -> comp_idL.
|
||||
Qed.
|
||||
Lemma UF_id : Unfold << Fold =-= Id.
|
||||
Proof.
|
||||
apply tset_trans with (y:=colimitExists ECoLimit ECoLimit).
|
||||
- apply (@colimitUnique _ ECoLimit ECoLimit). move => i. unfold Fold. unfold Unfold.
|
||||
simpl. rewrite -> lub_comp_both. rewrite -> lub_comp_left. apply Ole_antisym.
|
||||
* apply: (Ole_trans _ (le_lub _ i)).
|
||||
apply Oeq_le. simpl. do 2 rewrite -> ccomp_eq. simpl.
|
||||
do 3 rewrite comp_assoc. rewrite <- (comp_assoc (Embeddings i)).
|
||||
rewrite (emp i i). rewrite t_nn_ID. rewrite comp_idR.
|
||||
rewrite <- (comp_assoc (morph F _ _ _ _ _:M _ _)).
|
||||
rewrite -> (morph_comp F (Embeddings i) (Projections i)(Projections i) (Embeddings i)).
|
||||
rewrite -> (morph_eq_compat F (tset_trans (emp i i) (t_nn_ID _ _)) (tset_trans (emp i i) (t_nn_ID _ _))).
|
||||
rewrite -> morph_id. rewrite comp_idR. rewrite <- (comp_assoc (Injection i)).
|
||||
rewrite -> (proj1 (teppair DTower i)). by rewrite comp_idR.
|
||||
* rewrite -> (lub_lift_left _ i). apply: lub_le => k. simpl.
|
||||
apply Oeq_le. do 2 rewrite ccomp_eq.
|
||||
do 3 rewrite comp_assoc. rewrite <- (comp_assoc (Embeddings (i+k))).
|
||||
rewrite (emp _ _). rewrite t_nn_ID. rewrite comp_idR.
|
||||
rewrite <- (comp_assoc (morph F _ _ _ _ _:M _ _)).
|
||||
rewrite -> (morph_comp F (Embeddings (i + k)) (Projections (i + k)) (Projections i) (Embeddings i)).
|
||||
rewrite -> (morph_eq_compat F (emp i (i+k)) (emp (i+k) i)).
|
||||
rewrite -> morph_tnm. rewrite <- comp_assoc. rewrite <- (t_nmEmbedding DTower i (i + k).+1).
|
||||
rewrite <- comp_assoc. rewrite <- (t_nmProjection DTower i (i+k)). rewrite <- comp_assoc.
|
||||
rewrite <- (t_nmEmbedding2 DTower (leq_addr k i)). rewrite t_nmEmbedding. rewrite comp_assoc.
|
||||
rewrite <- morph_tnm.
|
||||
rewrite -> (morph_comp F (Projections (i+k)) (Embeddings (i+k)) (t_nm DTower (i+k) i) (t_nm DTower i (i+k))).
|
||||
by rewrite <- (morph_eq_compat F (@coneCom_l DTower (L DTower) _ _ (leq_addr k i)) (@coconeCom_l DTower DCoCone _ _ (leq_addr k i))).
|
||||
- apply tset_sym. apply (@colimitUnique _ ECoLimit ECoLimit). move => n. by rewrite -> comp_idL.
|
||||
Qed.
|
||||
(*=End *)
|
||||
|
||||
Add Parametric Morphism C X Y Z W : (@morph C X Y Z W)
|
||||
with signature (@Ole ((cppoMorph Y X) * (cppoMorph Z W))) ==> (@Ole _)
|
||||
as morph_le_compat.
|
||||
case => x x'. case => y y' e.
|
||||
by apply: fmonotonic.
|
||||
Qed.
|
||||
|
||||
Lemma delta_mon : monotonic (fun e => Fold << morph F _ _ _ _ (e,e) << Unfold).
|
||||
move => e e' X. simpl.
|
||||
by apply: (comp_le_compat (comp_le_compat (Ole_refl _) (morph_le_compat F (pair_le_compat X X))) (Ole_refl _)).
|
||||
Qed.
|
||||
|
||||
Definition deltat : ordCatType (cppoMorph DInf DInf) (cppoMorph DInf DInf) := Eval hnf in mk_fmono delta_mon.
|
||||
|
||||
Lemma comp_lub_eq (D0 D1 D2 :M) (f:D1 =-> D2) (c:natO =-> cppoMorph D0 D1) :
|
||||
f << lub c =-= lub ((exp_fun (ccomp D0 D1 D2) f : ordCatType _ _) << c).
|
||||
have a:=@lub_comp_eq _ _ ( (exp_fun (ccomp _ _ _) f)).
|
||||
specialize (a D0 c). simpl in a. rewrite <- a. by rewrite ccomp_eq.
|
||||
Qed.
|
||||
|
||||
Lemma lub_comp (D0 D1 D2 : M) (f:D0 =-> D1) (c:natO =-> cppoMorph D1 D2) :
|
||||
(lub c : M _ _) << f =-= lub ((exp_fun (ccomp D0 D1 D2 << <|pi2,pi1|>) f : ordCatType _ _) << c).
|
||||
have a:=@lub_comp_eq _ _ ( (exp_fun (ccomp _ _ _ << <|pi2,pi1|>) f)).
|
||||
specialize (a D2 c). rewrite <- a. simpl. by rewrite ccomp_eq.
|
||||
Qed.
|
||||
|
||||
Lemma delta_cont : continuous deltat.
|
||||
move => c. simpl. have S:=fmon_stable (morph F _ _ _ _).
|
||||
specialize (S _ _ _ _ _ _ (pair_lub c c)).
|
||||
rewrite -> (@lub_comp_eq _ _ ( (morph F _ _ _ _)) (<|c:natO =-> cppoMorph _ _,c|>)) in S.
|
||||
rewrite -> S. clear S. rewrite comp_lub_eq. rewrite lub_comp.
|
||||
apply: lub_le_compat => n. simpl. rewrite -> ccomp_eq. by rewrite ccomp_eq.
|
||||
Qed.
|
||||
|
||||
Definition delta : (cppoMorph DInf DInf) =-> (cppoMorph DInf DInf) :=
|
||||
Eval hnf in mk_fcont delta_cont.
|
||||
(*=Delta *)
|
||||
Lemma delta_simpl e : delta e = Fold << morph F _ _ _ _ (e,e) << Unfold.
|
||||
(*CLEAR*)
|
||||
by [].
|
||||
Qed. (*CLEARED*)
|
||||
(*=End *)
|
||||
|
||||
Require Import PredomFix.
|
||||
|
||||
Definition retract (C:catType) (A B : C) (f: A =-> B) (g : B =-> A) := f << g =-= Id.
|
||||
|
||||
Lemma retract_eq (C:catType) (X Y Z : C) (f:X =-> Y) g h (k:Z =-> X) : retract g f -> h =-= f << k -> g << h =-= k.
|
||||
move => R A.
|
||||
have B:=comp_eq_compat (tset_refl g) A. rewrite -> comp_assoc in B. rewrite -> R in B.
|
||||
by rewrite -> comp_idL in B.
|
||||
Qed.
|
||||
|
||||
Lemma delta_iter_id k : iter delta k <= (@Category.id M _:cppoMorph _ _).
|
||||
elim:k ; first by apply: leastP.
|
||||
move => n IH. simpl.
|
||||
rewrite -> (comp_le_compat (comp_le_compat (Ole_refl (Fold:cppoMorph _ _))
|
||||
(morph_le_compat F (pair_le_compat IH IH))) (Ole_refl (Unfold:cppoMorph _ _))).
|
||||
rewrite -> morph_id. rewrite comp_idR. by rewrite FU_id.
|
||||
Qed.
|
||||
|
||||
Lemma delta_id_id : delta (@Category.id M _:cppoMorph DInf DInf) =-= @Category.id M _.
|
||||
rewrite delta_simpl. rewrite morph_id. rewrite comp_idR. by rewrite FU_id.
|
||||
Qed.
|
||||
|
||||
(*=FIXDelta *)
|
||||
Lemma id_min : (FIXP delta : M _ _) =-= Id.
|
||||
(*=End *)
|
||||
rewrite <- EP_id.
|
||||
apply: (tset_trans (@limitUnique DTower (L DTower) (L DTower) _ _) _).
|
||||
- move => n. simpl. unfold fixp. rewrite -> lub_comp_right. apply Ole_antisym.
|
||||
+ rewrite <- (le_lub _ n). simpl. rewrite ccomp_eq.
|
||||
elim: n. simpl. apply Oeq_le. have a:= terminal_unique. specialize (a M). by apply: a.
|
||||
move => n IH. simpl. unfold Unfold. unfold Fold. simpl.
|
||||
rewrite <- (comp_le_compat (Ole_refl (Projections n.+1:cppoMorph _ _)) (comp_le_compat (comp_le_compat (le_lub (chainFPE DCoCone) n.+1) (Ole_refl (morph F _ _ _ _ (iter_ delta n, iter_ delta n)))) (le_lub (chainPEc ECoCone) n.+1))).
|
||||
simpl. repeat rewrite comp_assoc. rewrite -> (emp n.+1 n.+1). rewrite t_nn_ID. rewrite comp_idL.
|
||||
rewrite morph_comp. rewrite -> (morph_eq_compat F (mcoconeCom (DCoCone) n) (mconeCom (L DTower) n)).
|
||||
rewrite morph_comp. rewrite <- (comp_assoc (morph F _ _ _ _ (Projection n, Injection n):M _ _)).
|
||||
rewrite morph_comp. rewrite -> (morph_eq_compat F (mconeCom (L DTower) n) (mcoconeCom DCoCone n)).
|
||||
rewrite morph_comp.
|
||||
have a:= (Ole_trans _ (comp_le_compat (morph_le_compat F (pair_le_compat ((Oeq_le (tset_sym (comp_assoc _ (iter_ delta n:M _ _) _)))) (Ole_refl _))) (Ole_refl _))). apply: a.
|
||||
have aa:= (comp_le_compat (morph_le_compat F (pair_le_compat (comp_le_compat IH (Ole_refl _)) (comp_le_compat IH (Ole_refl _)))) (Ole_refl _)). rewrite <- aa. clear aa.
|
||||
rewrite (emp n n). rewrite t_nn_ID. rewrite morph_id. by rewrite comp_idL.
|
||||
+ apply: lub_le => k. simpl.
|
||||
rewrite ccomp_eq.
|
||||
rewrite -> (comp_le_compat (Ole_refl _) (delta_iter_id k)). by rewrite comp_idR.
|
||||
- apply tset_sym. apply limitUnique. move => n.
|
||||
rewrite EP_id. by rewrite comp_idR.
|
||||
Qed.
|
||||
|
||||
End RecursiveDomains.
|
||||
|
||||
End Msolution.
|
||||
|
||||
|
||||
|
||||
|
248
papers/domains-2012/PredomSum.v
Normal file
248
papers/domains-2012/PredomSum.v
Normal file
|
@ -0,0 +1,248 @@
|
|||
(**********************************************************************************
|
||||
* PredomSum.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Definition of co-products and associated lemmas
|
||||
==========================================================================*)
|
||||
Require Import PredomCore.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Order-theoretic definitions
|
||||
==========================================================================*)
|
||||
|
||||
Section OrdSum.
|
||||
|
||||
Variable O1 O2 : ordType.
|
||||
|
||||
Lemma sum_ordAxiom : PreOrd.axiom (fun (x y:O1+O2) => match x, y with | inl x, inl y => x <= y | inr x, inr y => x <= y | _,_ => False end).
|
||||
case => x0. split ; first by []. by case => x1 ; case => x2 ; try apply Ole_trans.
|
||||
split ; first by []. by case => x1 ; case => x2 ; try apply Ole_trans.
|
||||
Qed.
|
||||
|
||||
Canonical Structure sum_ordMixin := OrdMixin sum_ordAxiom.
|
||||
Canonical Structure sum_ordType := Eval hnf in OrdType sum_ordMixin.
|
||||
|
||||
Lemma inl_mono : monotonic (@inl O1 O2).
|
||||
by move => x y l.
|
||||
Qed.
|
||||
|
||||
Definition Inl : O1 =-> sum_ordType := Eval hnf in mk_fmono inl_mono.
|
||||
|
||||
Lemma inr_mono : monotonic (@inr O1 O2).
|
||||
by move => x y l.
|
||||
Qed.
|
||||
|
||||
Definition Inr : O2 =-> sum_ordType := Eval hnf in mk_fmono inr_mono.
|
||||
|
||||
Lemma sum_fun_mono (O3:ordType) (f:O1 =-> O3) (g:O2 =-> O3) : monotonic (fun x => match x with inl x => f x | inr y => g y end).
|
||||
by case => x ; case => y ; try done ; move => X ; apply fmonotonic.
|
||||
Qed.
|
||||
|
||||
Definition Osum_fun (O3:ordType) (f:O1 =-> O3) (g:O2 =-> O3) := Eval hnf in mk_fmono (sum_fun_mono f g).
|
||||
|
||||
End OrdSum.
|
||||
|
||||
Lemma sumOrdAxiom : @CatSum.axiom _ (sum_ordType) (@Inl) (@Inr) (@Osum_fun).
|
||||
move => O0 O1 O2 f g. split ; first by split ; apply: fmon_eq_intro.
|
||||
move => m [A B]. apply: fmon_eq_intro. case => x.
|
||||
- by apply (fmon_eq_elim A x).
|
||||
- by apply (fmon_eq_elim B x).
|
||||
Qed.
|
||||
|
||||
Canonical Structure sum_ordCatMixin := sumCatMixin sumOrdAxiom.
|
||||
Canonical Structure sum_ordCatType := Eval hnf in sumCatType sum_ordCatMixin.
|
||||
|
||||
Add Parametric Morphism (D E F:ordType) : (@Osum_fun D E F)
|
||||
with signature (@Ole _ : (D =-> F) -> (D =-> F) -> Prop) ++> (@Ole _ : (E =-> F) -> (E =-> F) -> Prop) ++>
|
||||
(@Ole (sum_ordType _ _ -=> _) : ((D + E) =-> F) -> (_ =-> _) -> Prop)
|
||||
as Osum_fun_le_compat.
|
||||
intros f f' lf g g' lg.
|
||||
simpl. intros x ; case x ; clear x ; auto.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D E F:ordType) : (@Osum_fun D E F)
|
||||
with signature (@tset_eq _ : (D =-> F) -> (D =-> F) -> Prop ) ==>
|
||||
(@tset_eq _ : (E =-> F) -> (E =-> F) -> Prop) ==> (@tset_eq ((sum_ordType D E) -=> F))
|
||||
as Osum_fun_eq_compat.
|
||||
intros f f' lf g g' lg.
|
||||
simpl. apply fmon_eq_intro.
|
||||
destruct lg. destruct lf.
|
||||
intros x ; split ; case x ; clear x ; auto.
|
||||
Qed.
|
||||
|
||||
Section CPOSum.
|
||||
|
||||
Variable D1 D2 : cpoType.
|
||||
|
||||
Lemma sum_left_mono (c:natO =-> sum_ordType D1 D2) (ex:{d| c O = inl D2 d}) :
|
||||
monotonic (fun x => match c x with | inl y => y | inr _ => projT1 ex end).
|
||||
case:ex => x X. simpl.
|
||||
move => z z' l. simpl.
|
||||
have H:= fmonotonic c (leq0n z). have HH:=fmonotonic c (leq0n z').
|
||||
rewrite -> X in H,HH.
|
||||
have A:=fmonotonic c l. case: (c z) H A ; last by [].
|
||||
move => x0 _. by case: (c z') HH.
|
||||
Qed.
|
||||
|
||||
Definition sum_left (c:natO =-> sum_ordType D1 D2) (ex:{d| c O = inl D2 d}) : natO =-> D1 :=
|
||||
Eval hnf in mk_fmono (sum_left_mono ex).
|
||||
|
||||
Lemma sum_right_mono (c:natO =-> sum_ordType D1 D2) (ex:{d| c O = inr D1 d}) :
|
||||
monotonic (fun x => match c x with | inr y => y | inl _ => projT1 ex end).
|
||||
case:ex => x X. simpl.
|
||||
move => z z' l. simpl.
|
||||
have H:= fmonotonic c (leq0n z). have HH:=fmonotonic c (leq0n z').
|
||||
rewrite -> X in H,HH.
|
||||
have A:=fmonotonic c l. case: (c z) H A ; first by [].
|
||||
move => x0 _. by case: (c z') HH.
|
||||
Qed.
|
||||
|
||||
Definition sum_right (c:natO =-> sum_ordType D1 D2) (ex:{d| c O = inr D1 d}) : natO =-> D2 :=
|
||||
Eval hnf in mk_fmono (sum_right_mono ex).
|
||||
|
||||
Definition SuminjProof (c:natO =-> sum_ordType D1 D2) (n:nat) :
|
||||
{d | c n = inl D2 d} + {d | c n = inr D1 d}.
|
||||
case: (c n) => cn.
|
||||
- left. by exists cn.
|
||||
- right. by exists cn.
|
||||
Defined.
|
||||
|
||||
Definition sum_lub (c: natO =-> sum_ordType D1 D2) : sum_ordType D1 D2 :=
|
||||
match (SuminjProof c 0) with
|
||||
| inl X => (inl D2 (lub (sum_left X)))
|
||||
| inr X => (inr D1 (lub (sum_right X)))
|
||||
end.
|
||||
|
||||
(*==========================================================================
|
||||
Domain-theoretic definitions
|
||||
==========================================================================*)
|
||||
|
||||
Lemma Dsum_ub (c : natO =-> sum_ordType D1 D2) (n : nat) : c n <= sum_lub c.
|
||||
have L:=fmonotonic c (leq0n n).
|
||||
unfold sum_lub. case: (SuminjProof c 0) ; case => c0 e ; rewrite e in L.
|
||||
- move: L. case_eq (c n) ; last by []. move => cn e' L. apply: (Ole_trans _ (le_lub _ n)).
|
||||
simpl. by rewrite e'.
|
||||
- move: L. case_eq (c n) ; first by []. move => cn e' L. apply: (Ole_trans _ (le_lub _ n)).
|
||||
simpl. by rewrite e'.
|
||||
Qed.
|
||||
|
||||
Lemma Dsum_lub (c : natO =-> sum_ordType D1 D2) (x : sum_ordType D1 D2) :
|
||||
(forall n : natO, c n <= x) -> sum_lub c <= x.
|
||||
case x ; clear x ; intros x cn. unfold sum_lub.
|
||||
case (SuminjProof c 0).
|
||||
- intros s. assert (lub (sum_left s) <= x).
|
||||
apply (lub_le). intros n. specialize (cn n). destruct s as [d c0].
|
||||
simpl. case_eq (c n).
|
||||
+ intros dn dnl. rewrite dnl in cn. by auto.
|
||||
+ intros dn cnr. rewrite cnr in cn. by inversion cn.
|
||||
by auto.
|
||||
- intros s. destruct s. simpl. specialize (cn 0). rewrite e in cn. by inversion cn.
|
||||
- unfold sum_lub. case (SuminjProof c 0) ; first case.
|
||||
+ move => y e. simpl. specialize (cn 0). rewrite e in cn. by inversion cn.
|
||||
+ intros s. assert (lub (sum_right s) <= x).
|
||||
* apply lub_le. intros n. specialize (cn n). destruct s as [d c0].
|
||||
simpl. case_eq (c n).
|
||||
- intros dl cnl. rewrite cnl in cn. by inversion cn.
|
||||
- intros dl dll. rewrite dll in cn. by auto.
|
||||
by auto.
|
||||
Qed.
|
||||
|
||||
Lemma sum_cpoAxiom : CPO.axiom sum_lub.
|
||||
move => c s n.
|
||||
split ; first by apply Dsum_ub.
|
||||
by apply Dsum_lub.
|
||||
Qed.
|
||||
|
||||
Canonical Structure sum_cpoMixin := CpoMixin sum_cpoAxiom.
|
||||
Canonical Structure sum_cpoType := Eval hnf in CpoType sum_cpoMixin.
|
||||
|
||||
(*
|
||||
(** Printing +c+ %\ensuremath{+}% *)
|
||||
Infix "+c+" := Dsum (at level 28, left associativity).*)
|
||||
|
||||
Variable D3:cpoType.
|
||||
|
||||
Lemma SUM_fun_cont (f:D1 =-> D3) (g:D2 =-> D3) : @continuous sum_cpoType _ (Osum_fun f g).
|
||||
intros c. simpl.
|
||||
case_eq (lub c).
|
||||
- move => ll sl. unfold lub in sl. simpl in sl. unfold sum_lub in sl. destruct (SuminjProof c 0) ; last by [].
|
||||
case: sl => e. rewrite <- e. clear ll e. rewrite -> (fcontinuous f (sum_left s)).
|
||||
apply lub_le_compat => n. simpl. case: s => d s. simpl. move: (fmonotonic c (leq0n n)).
|
||||
case: (c 0) s ; last by []. move => s e. rewrite -> e ; clear s e. by case: (c n).
|
||||
- move => ll sl. unfold lub in sl. simpl in sl. unfold sum_lub in sl. destruct (SuminjProof c 0) ; first by [].
|
||||
case: sl => e. rewrite <- e. clear ll e. rewrite -> (fcontinuous g (sum_right s)).
|
||||
apply lub_le_compat => n. simpl. case: s => d s. simpl. move: (fmonotonic c (leq0n n)).
|
||||
case: (c 0) s ; first by []. move => s e. rewrite -> e ; clear s e. by case: (c n).
|
||||
Qed.
|
||||
|
||||
Definition SUM_funX (f:D1 =-> D3) (g:D2 =-> D3) : sum_cpoType =-> D3 := Eval hnf in mk_fcont (SUM_fun_cont f g).
|
||||
Definition SUM_fun f g := locked (SUM_funX f g).
|
||||
|
||||
Lemma SUM_fun_simpl f g x : SUM_fun f g x =-= match x with inl x => f x | inr x => g x end.
|
||||
unlock SUM_fun. simpl. by case: x.
|
||||
Qed.
|
||||
|
||||
(** printing [| %\ensuremath{[} *)
|
||||
(** printing |] %\ensuremath{]} *)
|
||||
(* Notation "'[|' f , g '|]'" := (SUM_fun f g) (at level 30).*)
|
||||
|
||||
Lemma Inl_cont : continuous (Inl D1 D2).
|
||||
move => c. simpl. by apply: lub_le_compat => n.
|
||||
Qed.
|
||||
|
||||
Definition INL : D1 =-> sum_cpoType := Eval hnf in mk_fcont Inl_cont.
|
||||
|
||||
Lemma Inr_cont : continuous (Inr D1 D2).
|
||||
move => c. simpl. by apply: lub_le_compat => n.
|
||||
Qed.
|
||||
|
||||
Definition INR : D2 =-> sum_cpoType := Eval hnf in mk_fcont Inr_cont.
|
||||
|
||||
End CPOSum.
|
||||
|
||||
Lemma sumCpoAxiom : @CatSum.axiom _ sum_cpoType (@INL) (@INR) (@SUM_fun).
|
||||
move => D0 D1 D2 f g. split. unlock SUM_fun. apply: (proj1 (@sumOrdAxiom D0 D1 D2 f g)).
|
||||
move => m. unlock SUM_fun. apply: (proj2 (@sumOrdAxiom D0 D1 D2 f g)).
|
||||
Qed.
|
||||
|
||||
Canonical Structure cpoSumCatMixin := sumCatMixin sumCpoAxiom.
|
||||
Canonical Structure cpoSumCatType := Eval hnf in sumCatType cpoSumCatMixin.
|
||||
|
||||
Lemma SUM_fun_simplx (X Y Z : cpoType) (f:X =-> Z) (g : Y =-> Z) x : [|f,g|] x =-= match x with inl x => f x | inr x => g x end.
|
||||
case:x => x ; unfold sum_fun ; simpl ; by rewrite SUM_fun_simpl.
|
||||
Qed.
|
||||
|
||||
Definition getmorph (D E : cpoType) (f:D -=> E) : D =-> E := ev << <| const _ f, Id |>.
|
||||
|
||||
Lemma SUM_Fun_mon (D E F:cpoType) : monotonic (fun (P:(D -=> F) * (E -=> F)) => [|getmorph (fst P), getmorph (snd P)|]).
|
||||
case => f g. case => f' g'. case ; simpl => l l' x.
|
||||
unfold sum_fun. simpl. do 2 rewrite SUM_fun_simpl.
|
||||
by case x ; [apply l | apply l'].
|
||||
Qed.
|
||||
|
||||
Definition SUM_Funm (D E F:cpoType) : fmono_ordType ((D -=> F) * (E -=> F)) (D + E -=> F) :=
|
||||
Eval hnf in mk_fmono (@SUM_Fun_mon D E F).
|
||||
|
||||
Lemma SUM_Fon_cont (D E F:cpoType) : continuous (@SUM_Funm D E F).
|
||||
move => c. simpl. unfold getmorph. unfold sum_fun. simpl.
|
||||
by case => x ; simpl ; rewrite SUM_fun_simpl ; apply: lub_le_compat => n ; simpl ; rewrite SUM_fun_simpl.
|
||||
Qed.
|
||||
|
||||
Definition SUM_Fun (D E F:cpoType) : ((D -=> F) * (E -=> F)) =-> (D + E -=> F) := Eval hnf in mk_fcont (@SUM_Fon_cont D E F).
|
||||
Implicit Arguments SUM_Fun [D E F].
|
||||
|
||||
Lemma SUM_Fun_simpl (D E F:cpoType) (f:D =-> F) (g:E =-> F) : SUM_Fun (f,g) =-= SUM_fun f g.
|
||||
apply: fmon_eq_intro => x. simpl. unfold sum_fun. simpl. by do 2 rewrite SUM_fun_simpl.
|
||||
Qed.
|
||||
|
||||
|
1000
papers/domains-2012/mpremet.v
Normal file
1000
papers/domains-2012/mpremet.v
Normal file
File diff suppressed because it is too large
Load diff
390
papers/domains-2012/typedadequacy.v
Normal file
390
papers/domains-2012/typedadequacy.v
Normal file
|
@ -0,0 +1,390 @@
|
|||
(**********************************************************************************
|
||||
* typedadequacy.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
|
||||
(* Require Import utility.*)
|
||||
Require Import PredomAll.
|
||||
Require Import typedlambda.
|
||||
Require Import typedopsem.
|
||||
Require Import typeddensem.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Logical relation between semantics and terms, used to prove adequacy
|
||||
See Winskel Chapter 11, Section 11.4.
|
||||
==========================================================================*)
|
||||
|
||||
(** printing senv %\ensuremath{\rho}% *)
|
||||
(** printing d1 %\ensuremath{d_1}% *)
|
||||
|
||||
(* Lift a value-relation to a computation-relation *)
|
||||
(*=liftRel *)
|
||||
Definition liftRel t (R : SemTy t -> CValue t -> Prop) :=
|
||||
fun d e => forall d', d =-= Val d' -> exists v, e =>> v /\ R d' v. (*CLEAR*)
|
||||
(* = End *)
|
||||
|
||||
(* Logical relation between semantics and closed value terms *)
|
||||
Unset Implicit Arguments.
|
||||
|
||||
(* = relValEnvExp *)
|
||||
(*CLEARED*)
|
||||
Fixpoint relVal t : SemTy t -> CValue t -> Prop :=
|
||||
match t with
|
||||
| Int => fun d v => v = TINT d
|
||||
| Bool => fun d v => v = TBOOL d
|
||||
| t1 --> t2 => fun d v => exists e, v = TFIX e /\
|
||||
forall d1 v1, relVal t1 d1 v1 -> liftRel (relVal t2) (d d1) (subExp [ v1, v ] e)
|
||||
| t1 ** t2 => fun d v => exists v1, exists v2,
|
||||
v = TPAIR v1 v2 /\ relVal t1 (fst d) v1 /\ relVal t2 (snd d) v2
|
||||
end. (*CLEAR*)
|
||||
|
||||
Fixpoint relEnv E : SemEnv E -> Sub E nil -> Prop :=
|
||||
match E with
|
||||
| nil => fun _ _ => True
|
||||
| t :: E => fun d s => relVal t (snd d) (hdMap s) /\ relEnv E (fst d) (tlMap s)
|
||||
end.
|
||||
(*CLEARED*)
|
||||
Definition relExp ty := liftRel (relVal ty).
|
||||
(*=End *)
|
||||
(* Computation relation, already inlined into relVal *)
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
|
||||
Lemma relVal_lower_lift :
|
||||
forall ty,
|
||||
(forall d d' (v:CValue ty), d <= d' -> relVal ty d' v -> relVal ty d v) ->
|
||||
(forall d d' (e:CExp ty), d <= d' -> relExp ty d' e -> relExp ty d e).
|
||||
rewrite /relExp /liftRel. move => t IH d d' e l R d0 eq.
|
||||
rewrite -> eq in l. clear d eq. case: (DLle_Val_exists_eq l) => d'v [e0 l1].
|
||||
specialize (R _ e0). case: R => v [ev rv]. exists v. split ; first by [].
|
||||
by apply (IH _ _ _ l1).
|
||||
Qed.
|
||||
|
||||
(* Lemma 11.12 (ii) from Winskel *)
|
||||
(*=relValLower *)
|
||||
Lemma relVal_lower: forall t d d' v, d <= d' -> relVal t d' v -> relVal t d v.
|
||||
(*CLEAR*)
|
||||
|
||||
elim.
|
||||
|
||||
(* Int *)
|
||||
simpl. move => d d' v. by case => e ; rewrite e.
|
||||
|
||||
(* Bool *)
|
||||
simpl. move => d d' v. by case: d ; case: d'.
|
||||
|
||||
(* Arrow *)
|
||||
move => t IH t' IH' d d' v l /=. case => b. case => e ; rewrite e. clear v e.
|
||||
move => C. exists b. split ; first by []. move => d1 v1 r1. specialize (C d1 v1 r1).
|
||||
apply (relVal_lower_lift IH' (l d1)). by apply C.
|
||||
|
||||
(* Pair *)
|
||||
move => t IH t' IH' p1 p2 v L. case => v0 ; case => v1 [e l]. rewrite e. clear v e.
|
||||
simpl. exists v0. exists v1. split ; first by [].
|
||||
specialize (IH _ _ v0 (proj1 L)). specialize (IH' _ _ v1 (proj2 L)). split ; [apply IH | apply IH'] ; apply l.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism ty (v:CValue ty) : (fun d => relVal ty d v)
|
||||
with signature (@tset_eq (SemTy ty)) ==> iff as relVal_eq_compat.
|
||||
intros x y xyeq.
|
||||
destruct xyeq as [xy1 xy2].
|
||||
split; apply relVal_lower; trivial.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism t (v:CExp t) : (fun d => relExp t d v)
|
||||
with signature (@tset_eq ((SemTy t)_BOT)) ==> iff as relExp_eq_compat.
|
||||
intros x y xyeq.
|
||||
destruct xyeq as [xy1 xy2].
|
||||
split; apply (relVal_lower_lift (relVal_lower (t:=t))); trivial.
|
||||
Qed.
|
||||
|
||||
(* ==========================================================================
|
||||
Admissibility
|
||||
========================================================================== *)
|
||||
|
||||
Lemma rel_admissible_lift : forall ty,
|
||||
(forall v, admissible (fun d => relVal ty d v)) ->
|
||||
(forall e, admissible (fun d => relExp ty d e)).
|
||||
rewrite /admissible /relExp /liftRel. move => t IH e c C d eq.
|
||||
case: (lubval eq) => k [dv [ck dvd]].
|
||||
case: (C _ _ ck) => v [ev rv].
|
||||
exists v. split ; first by [].
|
||||
destruct (DLvalgetchain ck) as [c' P].
|
||||
specialize (IH v c').
|
||||
have El: ((@lub (SemTy t) c') =-= d). apply: vinj. rewrite <- eq.
|
||||
apply tset_trans with (y:=eta (lub c')) ; first by [].
|
||||
rewrite -> (lub_comp_eq). rewrite (lub_lift_left c k). apply: lub_eq_compat.
|
||||
apply fmon_eq_intro => n. simpl. by rewrite <- P.
|
||||
|
||||
refine (relVal_lower (proj2 El) _).
|
||||
apply IH => n.
|
||||
specialize (P n). specialize (C _ _ P). case: C => v' [ev' rv'].
|
||||
rewrite (Determinacy ev ev'). by apply rv'.
|
||||
Qed.
|
||||
|
||||
(* Lemma 11.12 (iii) from Winskel *)
|
||||
|
||||
(*CLEARED*)Lemma rel_admissible: forall t v, admissible (fun d => relVal t d v).
|
||||
(*=End *)
|
||||
elim.
|
||||
|
||||
(* Int *)
|
||||
intros v c C. simpl. apply (C 0).
|
||||
|
||||
(* Bool *)
|
||||
intros v c C. simpl. specialize (C 0). simpl in C. have e:lub c = c 0.
|
||||
unfold lub. simpl. unfold sum_lub. simpl. case: (SuminjProof c 0).
|
||||
+ simpl. case. case. move => e. by rewrite {2} e.
|
||||
+ simpl. case. case. move => e. by rewrite {2} e.
|
||||
by rewrite e.
|
||||
|
||||
(* Arrow *)
|
||||
rewrite /admissible => t IH t' IH' v c C.
|
||||
simpl.
|
||||
assert (rv := C 0). simpl in rv. destruct rv as [b [tbeq C0]].
|
||||
exists b.
|
||||
split; first by [].
|
||||
intros d1 v1 R1.
|
||||
refine (rel_admissible_lift _ _).
|
||||
intros vv cc CC. apply IH'. apply CC.
|
||||
|
||||
intros n. unfold relExp. unfold liftRel.
|
||||
intros d' CC.
|
||||
assert (Cn := C n). simpl in Cn.
|
||||
destruct Cn as [bb [tteq Cn]]. rewrite tteq in tbeq C C0 Cn. rewrite tteq. clear v tteq.
|
||||
assert (b = bb) as bbeq by apply (TFIX_injective (sym_equal tbeq)).
|
||||
rewrite <- bbeq in C0, C, Cn. rewrite <- bbeq. clear tbeq bb bbeq.
|
||||
specialize (Cn _ _ R1).
|
||||
unfold liftRel in Cn. specialize (Cn _ CC). by apply Cn.
|
||||
|
||||
(* Pair *)
|
||||
unfold admissible.
|
||||
move => t IH t' IH' v c C. simpl. unfold prod_lub.
|
||||
generalize C.
|
||||
destruct (ClosedPair v) as [v1 [v2 veq]]. subst.
|
||||
intros _.
|
||||
exists v1. exists v2.
|
||||
split;first by [].
|
||||
specialize (IH v1). specialize (IH' v2).
|
||||
split.
|
||||
- apply: IH. move => n. specialize (C n). case: C => v3 [v4 [veq [C1 C2]]].
|
||||
destruct (TPAIR_injective veq) as [E1 E2]. subst. by apply C1.
|
||||
- apply: IH'. move => n. specialize (C n). case: C => v3 [v4 [veq [C1 C2]]].
|
||||
destruct (TPAIR_injective veq) as [E1 E2]. subst. by apply C2.
|
||||
Qed.
|
||||
|
||||
Lemma SemEnvCons : forall t E (env : SemEnv (t :: E)), exists d, exists ds, env =(ds, d).
|
||||
intros. dependent inversion env. exists s0. by exists s.
|
||||
Qed.
|
||||
|
||||
Lemma Discrete_injective : forall t (v1 v2:discrete_cpoType t), v1 =-= v2 -> v1 = v2.
|
||||
intros t v1 v2 eq.
|
||||
destruct eq. destruct H. destruct H0. trivial.
|
||||
Qed.
|
||||
|
||||
Lemma relEnv_ext env senv s s' : (forall x y, s x y = s' x y) -> relEnv env senv s -> relEnv env senv s'.
|
||||
elim:env senv s s' ; first by [].
|
||||
move => a E IH env s s' C r.
|
||||
simpl. simpl in r. destruct r as [rl rr].
|
||||
split. unfold hdMap. rewrite <- (C a (ZVAR E a)). by apply rl.
|
||||
unfold tlMap.
|
||||
refine (IH _ _ _ _ rr). intros x y. by specialize (C x (SVAR a y)).
|
||||
Qed.
|
||||
|
||||
(*=FT *)
|
||||
Theorem FundamentalTheorem E:
|
||||
(forall t v senv s, relEnv E senv s -> relVal t (SemVal v senv) (subVal s v)) /\
|
||||
(forall t e senv s, relEnv E senv s -> liftRel (relVal t) (SemExp e senv) (subExp s e)).
|
||||
(*=End *)
|
||||
Proof.
|
||||
move: E ; apply ExpVal_ind.
|
||||
|
||||
(* TINT *)
|
||||
by simpl; auto.
|
||||
|
||||
(* TBOOL *)
|
||||
simpl. move=> E. by case.
|
||||
|
||||
(* TVAR *)
|
||||
intros E t v env s Renv.
|
||||
induction v.
|
||||
destruct (SemEnvCons env) as [d [ds eq]]. subst. simpl. fold SemEnv in *.
|
||||
simpl in Renv. destruct Renv as [R1 _]. by trivial.
|
||||
destruct (SemEnvCons env) as [d [ds eq]]. subst. simpl. fold SemEnv in *.
|
||||
simpl in Renv. destruct Renv as [_ R2].
|
||||
specialize (IHv ds (tlMap s) R2). by auto.
|
||||
|
||||
(* TFIX *)
|
||||
intros E t1 t2 e IH env s R.
|
||||
simpl SemVal.
|
||||
rewrite substTFIX.
|
||||
apply (@fixp_ind _ ((exp_fun (exp_fun (SemExp e)) env)) _ (@rel_admissible (t1 --> t2) _)).
|
||||
simpl. exists (subExp (liftSub t1 (liftSub (t1 --> t2) s)) e) ; split ; auto.
|
||||
intros d1 v1 rv1. unfold liftRel. intros dd.
|
||||
intros CC. by destruct (PBot_incon_eq (Oeq_sym CC)).
|
||||
|
||||
intros f rf.
|
||||
exists (subExp (liftSub t1 (liftSub (t1 --> t2) s)) e). split ; first by [].
|
||||
intros d1 v1 rv1. unfold liftRel. intros dd fd.
|
||||
assert (SemExp e (env,f,d1) =-= Val dd) as se by auto. clear fd.
|
||||
rewrite <- (proj2 (substComposeSub _)).
|
||||
rewrite <- substTFIX.
|
||||
assert (relEnv ((t1 :: t1 --> t2 :: E)) (env,f,d1) (consMap v1 (consMap (subVal (env':=nil) s (TFIX e)) s))).
|
||||
simpl relEnv. split.
|
||||
unfold hdMap. simpl. by apply rv1. split.
|
||||
unfold tlMap. unfold hdMap. simpl. rewrite substTFIX. exists (subExp
|
||||
(liftSub t1 (liftSub (t1 --> t2) s)) e). split ; first by [].
|
||||
intros d2 v2 rv2.
|
||||
simpl in rf. destruct rf as [bb [bbeq Pb]].
|
||||
assert (bb = (subExp (liftSub t1 (liftSub (t1 --> t2) s)) e)) as beq.
|
||||
apply (TFIX_injective (sym_equal bbeq)).
|
||||
rewrite beq in bbeq, Pb. clear bb beq bbeq. specialize (Pb d2 v2 rv2). by apply Pb.
|
||||
|
||||
simpl. unfold tlMap. simpl.
|
||||
refine (relEnv_ext _ R). by intros x y ; auto.
|
||||
|
||||
specialize (IH _ _ H). unfold liftRel in IH. specialize (IH _ se).
|
||||
rewrite composeCons. rewrite composeCons.
|
||||
rewrite composeSubIdLeft. by apply IH.
|
||||
|
||||
(* TP *)
|
||||
intros E t1 t2 v1 IH1 v2 IH2 env s R.
|
||||
simpl. specialize (IH1 env s R). specialize (IH2 env s R).
|
||||
exists (subVal s v1). exists (subVal s v2). by split; auto.
|
||||
|
||||
(* TFST *)
|
||||
intros E t1 t2 v IH env s R.
|
||||
simpl. specialize (IH env s R).
|
||||
unfold liftRel. intros d1 eq.
|
||||
rewrite substTFST.
|
||||
destruct (ClosedPair (subVal s v)) as [v1 [v2 veq]].
|
||||
rewrite veq in IH. rewrite veq. clear veq.
|
||||
exists v1. split. apply: e_Fst. simpl relVal in IH.
|
||||
destruct IH as [v3 [v4 [veq [gl1 gl2]]]]. destruct (TPAIR_injective veq) as [v1eq v2eq]. subst. clear veq.
|
||||
assert (H := vinj eq).
|
||||
by apply (relVal_lower (proj2 H) gl1).
|
||||
|
||||
(* TSND *)
|
||||
intros E t1 t2 v IH env s R.
|
||||
simpl. fold SemTy. specialize (IH env s R).
|
||||
unfold liftRel. intros d2 eq.
|
||||
rewrite substTSND.
|
||||
destruct (ClosedPair (subVal s v)) as [v1 [v2 veq]].
|
||||
rewrite veq in IH. rewrite veq. clear veq.
|
||||
exists v2. split. apply e_Snd. simpl relVal in IH.
|
||||
destruct IH as [v3 [v4 [veq [gl1 gl2]]]]. destruct (TPAIR_injective veq) as [v1eq v2eq]. subst. clear veq.
|
||||
assert (H := vinj eq).
|
||||
by apply (relVal_lower (proj2 H) gl2).
|
||||
|
||||
(* TOP *)
|
||||
intros E op v1 IH1 v2 IH2 env s R.
|
||||
simpl. specialize (IH1 env s R). specialize (IH2 env s R).
|
||||
unfold liftRel. intros d eq.
|
||||
rewrite substTOP.
|
||||
destruct (ClosedInt (subVal s v1)) as [i1 eq1].
|
||||
destruct (ClosedInt (subVal s v2)) as [i2 eq2].
|
||||
rewrite -> eq1 in IH1. rewrite eq1. rewrite eq2 in IH2. rewrite eq2. clear eq1 eq2.
|
||||
exists (TINT (op i1 i2)).
|
||||
split. apply e_Op. simpl in IH1, IH2. case: IH1 => IH1. rewrite -> IH1.
|
||||
case: IH2 => e ; rewrite -> e. case: (vinj eq). simpl => ee. by rewrite <- ee.
|
||||
|
||||
(* TGT *)
|
||||
intros E v1 IH1 v2 IH2 env s R.
|
||||
simpl. specialize (IH1 env s R). specialize (IH2 env s R).
|
||||
unfold liftRel. intros d eq. have H:=vinj eq.
|
||||
rewrite substTGT.
|
||||
destruct (ClosedInt (subVal s v1)) as [i1 eq1].
|
||||
destruct (ClosedInt (subVal s v2)) as [i2 eq2].
|
||||
rewrite -> eq1 in IH1. rewrite eq1. rewrite eq2 in IH2. rewrite eq2. clear eq1 eq2.
|
||||
exists (TBOOL (i2 <= i1)%N).
|
||||
split. apply e_Gt. simpl in IH1,IH2. case: IH1 => e ; rewrite e. clear e. case: IH2 => e ; rewrite e. clear e.
|
||||
f_equal.
|
||||
have HH:(if (SemVal v2 env <= SemVal v1 env)%N then INL _ _ tt else INR _ _ tt) =-= d by apply H. clear H.
|
||||
case: (SemVal v2 env <= SemVal v1 env)%N HH ; simpl ; clear eq ; case: d ; simpl ; case ; by case.
|
||||
|
||||
(* VAL *)
|
||||
intros E t v IH env s R.
|
||||
simpl. specialize (IH env s R).
|
||||
unfold liftRel. intros d eq. have H:=vinj eq. clear eq.
|
||||
exists (subVal s v).
|
||||
split. apply e_Val.
|
||||
apply (relVal_lower (proj2 H) IH).
|
||||
|
||||
(* TLET *)
|
||||
intros E t1 t2 e1 IH1 e2 IH2 env s R.
|
||||
unfold liftRel. intros d2 d2eq.
|
||||
specialize (IH1 env s R). unfold liftRel in *.
|
||||
simpl in d2eq.
|
||||
|
||||
rewrite substTLET.
|
||||
destruct (KLEISLIR_ValVal d2eq) as [d1 [d1eq d2veq]]. clear d2eq.
|
||||
specialize (IH1 _ d1eq).
|
||||
destruct IH1 as [v1 [ev1 R1]].
|
||||
specialize (IH2 (env,d1) (consMap v1 s)).
|
||||
assert (RR : relEnv (t1::E) (env,d1) (consMap v1 s)).
|
||||
simpl relEnv. split ; first by [].
|
||||
by rewrite tlConsMap.
|
||||
specialize (IH2 RR d2 d2veq).
|
||||
destruct IH2 as [v2 [ev2 R2]].
|
||||
exists v2. split.
|
||||
refine (e_Let ev1 _).
|
||||
rewrite <- (proj2 (substComposeSub _)). rewrite composeCons. rewrite composeSubIdLeft. apply ev2. by assumption.
|
||||
|
||||
(* TAPP *)
|
||||
intros E t1 t2 v1 IH1 v2 IH2 env s R.
|
||||
specialize (IH1 env s R). specialize (IH2 env s R).
|
||||
unfold liftRel. intros d eq. rewrite substTAPP.
|
||||
destruct (ClosedFunction (subVal s v1)) as [e1 eq1].
|
||||
unfold relVal in IH1. fold relVal in IH1. destruct IH1 as [e1' [eq' R']].
|
||||
rewrite eq' in eq1.
|
||||
assert (H := TFIX_injective eq1).
|
||||
subst. clear eq1.
|
||||
specialize (R' (SemVal v2 env) (subVal s v2) IH2).
|
||||
clear IH2.
|
||||
unfold liftRel in R'. specialize (R' d eq).
|
||||
destruct R' as [v [ev Rv]].
|
||||
exists v. rewrite eq' in ev.
|
||||
assert (H := e_App ev).
|
||||
split. simpl. rewrite eq'. apply H.
|
||||
by apply Rv.
|
||||
|
||||
(* TIF *)
|
||||
intros E t v IH e1 IH1 e2 IH2 env s R.
|
||||
specialize (IH env s R). specialize (IH1 env s R). specialize (IH2 env s R).
|
||||
unfold liftRel in *. intros d eq.
|
||||
rewrite substTIF.
|
||||
destruct (ClosedBool (subVal s v)) as [b beq].
|
||||
rewrite beq in IH.
|
||||
unfold relVal in IH.
|
||||
inversion IH. clear IH.
|
||||
simpl subExp. rewrite beq.
|
||||
rewrite H0 in beq. rewrite H0. clear b H0. simpl in eq.
|
||||
case: (SemVal v env) eq beq ; case => eq beq.
|
||||
- specialize (IH1 _ eq). case: IH1 => v0 [S0 R0]. exists v0. split ; last by [].
|
||||
by apply (e_IfTrue _ S0).
|
||||
- specialize (IH2 _ eq). case: IH2 => v0 [S0 R0]. exists v0. split ; last by [].
|
||||
by apply (e_IfFalse _ S0).
|
||||
Qed.
|
||||
|
||||
(*=Adequacy *)
|
||||
Corollary Adequacy: forall t (e : CExp t) d, SemExp e tt =-= Val d -> exists v, e =>> v.
|
||||
(*=End *)
|
||||
Proof.
|
||||
intros t e d val.
|
||||
destruct (FundamentalTheorem nil) as [_ FT].
|
||||
specialize (FT t e tt (idSub nil)).
|
||||
have X: (relEnv nil tt (idSub nil)) by [].
|
||||
specialize (FT X). unfold liftRel in FT. specialize (FT d val).
|
||||
destruct FT as [v [ev _]]. exists v. by rewrite (proj2 (applyIdSub _)) in ev.
|
||||
Qed.
|
||||
|
114
papers/domains-2012/typeddensem.v
Normal file
114
papers/domains-2012/typeddensem.v
Normal file
|
@ -0,0 +1,114 @@
|
|||
(**********************************************************************************
|
||||
* typeddensem.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
Require Export typedlambda.
|
||||
Require Import PredomAll.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
|
||||
Open Scope C_scope.
|
||||
Open Scope D_scope.
|
||||
Open Scope S_scope.
|
||||
|
||||
(*==========================================================================
|
||||
Meaning of types and contexts
|
||||
==========================================================================*)
|
||||
|
||||
Delimit Scope Ty_scope with Ty.
|
||||
|
||||
Open Scope Ty_scope.
|
||||
|
||||
(*=SemTy *)
|
||||
Fixpoint SemTy t : cpoType :=
|
||||
match t with
|
||||
| Int => nat_cpoType
|
||||
| Bool => (One:cpoType) + One
|
||||
| t1 --> t2 => SemTy t1 -=> (SemTy t2)_BOT
|
||||
| t1 ** t2 => SemTy t1 * SemTy t2
|
||||
end.
|
||||
Fixpoint SemEnv E : cpoType :=
|
||||
match E with
|
||||
| nil => One
|
||||
| t :: E => SemEnv E * SemTy t
|
||||
end.
|
||||
(*=End *)
|
||||
|
||||
(* Lookup in an environment *)
|
||||
(*=SemVarExpVal *)
|
||||
Fixpoint SemVar E t (var : Var E t) : SemEnv E =-> SemTy t :=
|
||||
match var with
|
||||
| ZVAR _ _ => pi2
|
||||
| SVAR _ _ _ v => SemVar v << pi1
|
||||
end. (*CLEAR*)
|
||||
|
||||
Lemma SimpleB_mon (A B :Type) (C:ordType) (op : A -> B -> C:Type) : monotonic (fun p:discrete_ordType A * discrete_ordType B => op (fst p) (snd p)).
|
||||
case => x0 y0. case => x1 y1. case ; simpl. move => L L'.
|
||||
have e:x0 = x1 by []. have e':y0 = y1 by []. rewrite e. by rewrite e'.
|
||||
Qed.
|
||||
|
||||
Definition SimpleBOpm (A B :Type) (C:ordType) (op : A -> B -> C:Type) : discrete_ordType A * discrete_ordType B =-> C :=
|
||||
mk_fmono (SimpleB_mon op).
|
||||
|
||||
Lemma SimpleB_cont (A B:Type) (C:cpoType) (op : A -> B -> C:Type) :
|
||||
@continuous (discrete_cpoType A * discrete_cpoType B) C (SimpleBOpm op).
|
||||
move => c. simpl. apply: (Ole_trans _ (le_lub _ 0)). simpl.
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition SimpleBOp (A B:Type) (C:cpoType) (op : A -> B -> C:Type) : discrete_cpoType A * discrete_cpoType B =-> C :=
|
||||
Eval hnf in mk_fcont (SimpleB_cont op).
|
||||
|
||||
Lemma choose_cont (Y Z: cpoType) (a:Z =-> Y) (b : Z =-> Y) (c:Z =-> (One:cpoType) + One) :
|
||||
forall x,
|
||||
(fun y => match c y with inl y' => a x | inr y' => b x end) x =-=
|
||||
(ev << <| SUM_Fun << <| exp_fun ((uncurry (const One a)) << SWAP),
|
||||
exp_fun ((uncurry (const One b)) << SWAP) |>, c|>) x.
|
||||
move => x. simpl. rewrite SUM_fun_simplx. by case_eq (c x) ; case => e.
|
||||
Qed.
|
||||
|
||||
Definition choose (Y Z: cpoType) (a:Z =-> Y) (b : Z =-> Y) (c:Z =-> (One:cpoType) + One) :=
|
||||
Eval hnf in gen_cont (choose_cont a b c).
|
||||
|
||||
Lemma choose_comp (Y Z W : cpoType) (a:Z =-> Y) (b : Z =-> Y) (c:Z =-> (One:cpoType) + One)
|
||||
(h : W =-> Z) : choose a b c << h =-= choose (a << h) (b << h) (c << h).
|
||||
by apply: fmon_eq_intro => x.
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (A B:cpoType) : (@choose A B)
|
||||
with signature (@tset_eq _) ==> (@tset_eq _) ==> (@tset_eq _) ==> (@tset_eq _)
|
||||
as choose_eq_compat.
|
||||
move => f f' e g g' e' x x' ex.
|
||||
apply: fmon_eq_intro => y. simpl. move: (fmon_eq_elim ex y). clear ex. unfold FCont.fmono. simpl.
|
||||
case: (x y) ; case: (x' y) ; case ; case ; (try by case) ; move => _. by rewrite e. by rewrite e'.
|
||||
Qed.
|
||||
|
||||
(* SemExpVal *)
|
||||
(*CLEARED*)
|
||||
Fixpoint SemExp E t (e : Exp E t) : SemEnv E =-> (SemTy t) _BOT :=
|
||||
match e with
|
||||
| TOP op v1 v2 => eta << SimpleBOp op << <| SemVal v1 , SemVal v2 |>
|
||||
| TGT v1 v2 => eta << SimpleBOp (fun x y => if leq x y then inl _ tt else inr _ tt)
|
||||
<< <| SemVal v2 , SemVal v1 |>
|
||||
| TAPP _ _ v1 v2 => ev << <| SemVal v1 , SemVal v2 |>
|
||||
| TVAL _ v => eta << SemVal v
|
||||
| TLET _ _ e1 e2 => KLEISLIR (SemExp e2) << <| Id , SemExp e1 |>
|
||||
| TIF _ v e1 e2 => choose (SemExp e1) (SemExp e2) (SemVal v)
|
||||
| TFST _ _ v => eta << pi1 << SemVal v
|
||||
| TSND _ _ v => eta << pi2 << SemVal v
|
||||
end with SemVal E t (v : Value E t) : SemEnv E =-> SemTy t :=
|
||||
match v with
|
||||
| TINT n => const _ n
|
||||
| TBOOL b => const _ (if b then inl _ tt else inr _ tt)
|
||||
| TVAR _ i => SemVar i
|
||||
| TFIX t1 t2 e => (FIXP : cpoCatType _ _) << exp_fun (exp_fun (SemExp e))
|
||||
| TPAIR _ _ v1 v2 => <| SemVal v1 , SemVal v2 |>
|
||||
end.
|
||||
(*=End *)
|
505
papers/domains-2012/typedlambda.v
Normal file
505
papers/domains-2012/typedlambda.v
Normal file
|
@ -0,0 +1,505 @@
|
|||
(**********************************************************************************
|
||||
* typedlambda.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Call-by-value simply-typed lambda-calculus with recursion, pairs, and arithmetic
|
||||
==========================================================================*)
|
||||
|
||||
Require Import Program.Equality.
|
||||
Require Export ssreflect ssrnat seq eqtype.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
Set Printing Implicit Defensive.
|
||||
|
||||
(*==========================================================================
|
||||
Types and contexts
|
||||
==========================================================================*)
|
||||
|
||||
(*=Ty *)
|
||||
Inductive Ty := Int | Bool | Arrow (ty1 t2 : Ty) | Prod (ty1 t2 : Ty).
|
||||
Infix " --> " := Arrow (at level 55, right associativity).
|
||||
Infix " ** " := Prod (at level 55).
|
||||
Definition Env := seq Ty.
|
||||
(*=End *)
|
||||
(* begin hide *)
|
||||
Notation "[ x , .. , y ]" := (cons x .. (cons y nil) ..) : list_scope.
|
||||
(* end hide *)
|
||||
|
||||
(*==========================================================================
|
||||
Typed terms in context
|
||||
==========================================================================*)
|
||||
(*=VarValueExp *)
|
||||
Inductive Var : Env -> Ty -> Type :=
|
||||
| ZVAR : forall E t, Var (t :: E) t
|
||||
| SVAR : forall E t t', Var E t -> Var (t' :: E) t.
|
||||
Inductive Value E : Ty -> Type :=
|
||||
| TINT : nat -> Value E Int
|
||||
| TBOOL : bool -> Value E Bool
|
||||
| TVAR :> forall t, Var E t -> Value E t
|
||||
| TFIX : forall t1 t2, Exp (t1 :: t1 --> t2 :: E) t2 -> Value E (t1 --> t2)
|
||||
| TPAIR : forall t1 t2, Value E t1 -> Value E t2 -> Value E (t1 ** t2)
|
||||
with Exp E : Ty -> Type :=
|
||||
| TFST : forall t1 t2, Value E (t1 ** t2) -> Exp E t1
|
||||
| TSND : forall t1 t2, Value E (t1 ** t2) -> Exp E t2
|
||||
| TOP : (nat -> nat -> nat) -> Value E Int -> Value E Int -> Exp E Int
|
||||
| TGT : Value E Int -> Value E Int -> Exp E Bool
|
||||
| TVAL : forall t, Value E t -> Exp E t
|
||||
| TLET : forall t1 t2, Exp E t1 -> Exp (t1 :: E) t2 -> Exp E t2
|
||||
| TAPP : forall t1 t2, Value E (t1 --> t2) -> Value E t1 -> Exp E t2
|
||||
| TIF : forall t, Value E Bool -> Exp E t -> Exp E t -> Exp E t.
|
||||
Definition CExp t := Exp nil t.
|
||||
Definition CValue t := Value nil t.
|
||||
(*=End *)
|
||||
Implicit Arguments TBOOL [E].
|
||||
Implicit Arguments TINT [E].
|
||||
|
||||
(* begin hide *)
|
||||
Scheme Val_rec2 := Induction for Value Sort Set
|
||||
with Exp_rec2 := Induction for Exp Sort Set.
|
||||
|
||||
Scheme Val_type2 := Induction for Value Sort Type
|
||||
with Exp_type2 := Induction for Exp Sort Type.
|
||||
|
||||
Scheme Val_ind2 := Induction for Value Sort Prop
|
||||
with Exp_ind2 := Induction for Exp Sort Prop.
|
||||
|
||||
Combined Scheme ExpVal_ind from Val_ind2, Exp_ind2.
|
||||
(* end hide *)
|
||||
|
||||
(*==========================================================================
|
||||
Variable-domain maps.
|
||||
By instantiating P with Var we get renamings.
|
||||
By instantiating P with Val we get substitutions.
|
||||
==========================================================================*)
|
||||
|
||||
Definition Map (P:Env -> Ty -> Type) E E' := forall t, Var E t -> P E' t.
|
||||
|
||||
(* Head, tail and cons *)
|
||||
Definition tlMap P E E' t (m:Map P (t::E) E') : Map P E E' := fun t' v => m t' (SVAR t v).
|
||||
Definition hdMap P E E' t (m:Map P (t::E) E') : P E' t := m t (ZVAR _ _).
|
||||
Implicit Arguments tlMap [P E E' t].
|
||||
Implicit Arguments hdMap [P E E' t].
|
||||
|
||||
|
||||
Definition cast P (E E' : Env) (t : Ty) (x:P E t) (p:E=E') : P E' t.
|
||||
intros.
|
||||
subst. exact x.
|
||||
Defined.
|
||||
|
||||
Program Definition consMap P E E' t (v:P E' t) (m:Map P E E') : Map P (t::E) E' :=
|
||||
fun t' (var:Var (t::E) t') =>
|
||||
match var with
|
||||
| ZVAR _ _ => v
|
||||
| SVAR _ _ _ var => m _ var
|
||||
end.
|
||||
|
||||
Implicit Arguments consMap [P E E' t].
|
||||
|
||||
Axiom MapExtensional : forall P E E' (r1 r2 : Map P E E'), (forall t var, r1 t var = r2 t var) -> r1 = r2.
|
||||
|
||||
Lemma hdConsMap : forall P env env' ty (v : P env' ty) (s : Map P env env'), hdMap (consMap v s) = v. Proof. auto. Qed.
|
||||
Lemma tlConsMap : forall P env env' ty (v : P env' ty) (s : Map P env env'), tlMap (consMap v s) = s. Proof. intros. apply MapExtensional. auto. Qed.
|
||||
|
||||
Lemma consMapInv : forall P env env' ty (m:Map P (ty :: env) env'), exists m', exists v, m = consMap v m'.
|
||||
intros. exists (tlMap m). exists (hdMap m).
|
||||
apply MapExtensional. dependent destruction var; auto.
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Package of operations used with a Map
|
||||
vr maps a Var into Var or Value (so is either the identity or TVAR)
|
||||
vl maps a Var or Value to a Value (so is either TVAR or the identity)
|
||||
wk weakens a Var or Value (so is either SVAR or renaming through SVAR on a value)
|
||||
==========================================================================*)
|
||||
Record MapOps (P : Env -> Ty -> Type) :=
|
||||
{
|
||||
vr : forall E ty, Var E ty -> P E ty;
|
||||
vl : forall E ty, P E ty -> Value E ty;
|
||||
wk : forall E ty ty', P E ty -> P (ty' :: E) ty
|
||||
}.
|
||||
|
||||
Section MapOps.
|
||||
|
||||
Variable P : Env -> Ty -> Type.
|
||||
Variable ops : MapOps P.
|
||||
|
||||
Program Definition lift env env' ty (m : Map P env env') : Map P (ty :: env) (ty :: env') :=
|
||||
fun ty' => fun var => match var with
|
||||
| ZVAR _ _ => vr ops (ZVAR _ _)
|
||||
| SVAR _ _ _ x => wk ops _ (m _ x)
|
||||
end.
|
||||
Implicit Arguments lift [env env'].
|
||||
|
||||
Definition shiftMap env env' ty (m : Map P env env') : Map P env (ty :: env') := fun ty' => fun var => wk ops _ (m ty' var).
|
||||
Implicit Arguments shiftMap [env env'].
|
||||
|
||||
Lemma shiftConsMap : forall env env' ty (m : Map P env env') (x : P env' ty) ty', shiftMap ty' (consMap x m) = consMap (wk ops _ x) (shiftMap ty' m).
|
||||
Proof.
|
||||
intros env env' ty m x ty'. apply MapExtensional.
|
||||
intros ty0 var0. dependent destruction var0; auto.
|
||||
Qed.
|
||||
|
||||
Fixpoint travVal env env' ty (v : Value env ty) : Map P env env' -> Value env' ty :=
|
||||
match v with
|
||||
| TINT i => fun m => TINT i
|
||||
| TBOOL b => fun m => TBOOL b
|
||||
| TVAR _ v => fun m => vl ops (m _ v)
|
||||
| TFIX _ _ e => fun m => TFIX (travExp e (lift _ (lift _ m)))
|
||||
| TPAIR _ _ e1 e2 => fun m => TPAIR (travVal e1 m) (travVal e2 m)
|
||||
end
|
||||
with travExp env env' ty (e : Exp env ty) : Map P env env' -> Exp env' ty :=
|
||||
match e with
|
||||
| TOP op e1 e2 => fun m => TOP op (travVal e1 m) (travVal e2 m)
|
||||
| TGT e1 e2 => fun m => TGT (travVal e1 m) (travVal e2 m)
|
||||
| TFST _ _ e => fun m => TFST (travVal e m)
|
||||
| TSND _ _ e => fun m => TSND (travVal e m)
|
||||
| TVAL _ v => fun m => TVAL (travVal v m)
|
||||
| TIF _ v e1 e2 => fun m => TIF (travVal v m) (travExp e1 m) (travExp e2 m)
|
||||
| TAPP _ _ e1 e2 => fun m => TAPP (travVal e1 m) (travVal e2 m)
|
||||
| TLET _ _ e1 e2 => fun m => TLET (travExp e1 m) (travExp e2 (lift _ m))
|
||||
end.
|
||||
|
||||
Definition mapVal env env' ty m v := @travVal env env' ty v m.
|
||||
Definition mapExp env env' ty m e := @travExp env env' ty e m.
|
||||
|
||||
Variable env env' : Env.
|
||||
Variable m : Map P env env'.
|
||||
Variable t1 t2 : Ty.
|
||||
|
||||
Lemma mapTVAR : forall (var : Var _ t1), mapVal m (TVAR var) = vl ops (m var). auto. Qed.
|
||||
Lemma mapTINT : forall n, mapVal m (TINT n) = TINT n. auto. Qed.
|
||||
Lemma mapTBOOL : forall n, mapVal m (TBOOL n) = TBOOL n. auto. Qed.
|
||||
Lemma mapTPAIR : forall (v1 : Value _ t1) (v2 : Value _ t2), mapVal m (TPAIR v1 v2) = TPAIR (mapVal m v1) (mapVal m v2). auto. Qed.
|
||||
Lemma mapTFST : forall (v : Value _ (t1 ** t2)), mapExp m (TFST v) = TFST (mapVal m v). auto. Qed.
|
||||
Lemma mapTSND : forall (v : Value _ (t1 ** t2)), mapExp m (TSND v) = TSND (mapVal m v). auto. Qed.
|
||||
Lemma mapTFIX : forall (e : Exp (t1 :: t1-->t2 :: _) t2), mapVal m (TFIX e) = TFIX (mapExp (lift _ (lift _ m)) e). auto. Qed.
|
||||
Lemma mapTOP : forall op v1 v2, mapExp m (TOP op v1 v2) = TOP op (mapVal m v1) (mapVal m v2). auto. Qed.
|
||||
Lemma mapTGT : forall v1 v2, mapExp m (TGT v1 v2) = TGT (mapVal m v1) (mapVal m v2). auto. Qed.
|
||||
Lemma mapTVAL : forall (v : Value _ t1), mapExp m (TVAL v) = TVAL (mapVal m v). auto. Qed.
|
||||
Lemma mapTLET : forall (e1 : Exp _ t1) (e2 : Exp _ t2), mapExp m (TLET e1 e2) = TLET (mapExp m e1) (mapExp (lift _ m) e2). auto. Qed.
|
||||
Lemma mapTIF : forall v (e1 e2 : Exp _ t1), mapExp m (TIF v e1 e2) = TIF (mapVal m v) (mapExp m e1) (mapExp m e2). auto. Qed.
|
||||
Lemma mapTAPP : forall (v1 : Value _ (t1 --> t2)) v2, mapExp m (TAPP v1 v2) = TAPP (mapVal m v1) (mapVal m v2). auto. Qed.
|
||||
|
||||
End MapOps.
|
||||
|
||||
Hint Rewrite mapTVAR mapTINT mapTBOOL mapTPAIR mapTFST mapTSND mapTFIX mapTOP mapTGT mapTVAL mapTLET mapTIF mapTAPP : mapHints.
|
||||
|
||||
Implicit Arguments lift [P env env'].
|
||||
Implicit Arguments shiftMap [P env env'].
|
||||
|
||||
(*==========================================================================
|
||||
Variable renamings: Map Var
|
||||
==========================================================================*)
|
||||
|
||||
Definition Ren := Map Var.
|
||||
Definition RenMapOps := (Build_MapOps (fun _ _ v => v) TVAR SVAR).
|
||||
|
||||
Definition renameVal := mapVal RenMapOps.
|
||||
Definition renameExp := mapExp RenMapOps.
|
||||
Definition liftRen := lift RenMapOps.
|
||||
Implicit Arguments liftRen [env env'].
|
||||
Definition shiftRen := shiftMap RenMapOps.
|
||||
Implicit Arguments shiftRen [env env'].
|
||||
|
||||
Section RenDefs.
|
||||
|
||||
Variable env env' : Env.
|
||||
Variable r : Ren env env'.
|
||||
Variable t1 t2 : Ty.
|
||||
|
||||
Lemma renameTVAR : forall (var : Var env t1), renameVal r (TVAR var) = TVAR (r var). auto. Qed.
|
||||
Lemma renameTINT : forall n, renameVal r (TINT n) = TINT n. auto. Qed.
|
||||
Lemma renameTBOOL : forall n, renameVal r (TBOOL n) = TBOOL n. auto. Qed.
|
||||
Lemma renameTPAIR : forall (v1 : Value _ t1) (v2 : Value _ t2), renameVal r (TPAIR v1 v2) = TPAIR (renameVal r v1) (renameVal r v2). auto. Qed.
|
||||
Lemma renameTFST : forall (v : Value _ (t1 ** t2)), renameExp r (TFST v) = TFST (renameVal r v). auto. Qed.
|
||||
Lemma renameTSND : forall (v : Value _ (t1 ** t2)), renameExp r (TSND v) = TSND (renameVal r v). auto. Qed.
|
||||
Lemma renameTFIX : forall (e : Exp (t1 :: t1-->t2 :: _) t2), renameVal r (TFIX e) = TFIX (renameExp (liftRen _ (liftRen _ r)) e). auto. Qed.
|
||||
Lemma renameTOP : forall op v1 v2, renameExp r (TOP op v1 v2) = TOP op (renameVal r v1) (renameVal r v2). auto. Qed.
|
||||
Lemma renameTGT : forall v1 v2, renameExp r (TGT v1 v2) = TGT (renameVal r v1) (renameVal r v2). auto. Qed.
|
||||
Lemma renameTVAL : forall (v : Value env t1), renameExp r (TVAL v) = TVAL (renameVal r v). auto. Qed.
|
||||
Lemma renameTLET : forall (e1 : Exp _ t1) (e2 : Exp _ t2), renameExp r (TLET e1 e2) = TLET (renameExp r e1) (renameExp (liftRen _ r) e2). auto. Qed.
|
||||
Lemma renameTIF : forall v (e1 e2 : Exp _ t1), renameExp r (TIF v e1 e2) = TIF (renameVal r v) (renameExp r e1) (renameExp r e2). auto. Qed.
|
||||
Lemma renameTAPP : forall (v1 : Value _ (t1-->t2)) v2, renameExp r (TAPP v1 v2) = TAPP (renameVal r v1) (renameVal r v2). auto. Qed.
|
||||
|
||||
End RenDefs.
|
||||
|
||||
Hint Rewrite renameTVAR renameTINT renameTBOOL renameTPAIR renameTFST renameTSND renameTFIX renameTOP renameTGT renameTVAL renameTLET renameTIF renameTAPP : renameHints.
|
||||
|
||||
Lemma LiftRenDef : forall env env' (r : Ren env' env) ty, liftRen _ r = consMap (ZVAR _ ty) (shiftRen _ r).
|
||||
intros. apply MapExtensional. auto. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Identity renaming
|
||||
==========================================================================*)
|
||||
|
||||
Definition idRen env : Ren env env := fun ty (var : Var env ty) => var.
|
||||
Implicit Arguments idRen [].
|
||||
|
||||
Lemma liftIdRen : forall E t, liftRen _ (idRen E) = idRen (t::E).
|
||||
intros. apply MapExtensional.
|
||||
dependent destruction var; auto.
|
||||
Qed.
|
||||
|
||||
Lemma applyIdRen env :
|
||||
(forall ty (v : Value env ty), renameVal (idRen env) v = v)
|
||||
/\ (forall ty (e : Exp env ty), renameExp (idRen env) e = e).
|
||||
move: env. apply ExpVal_ind;
|
||||
(intros; try (autorewrite with renameHints using try rewrite liftIdRen; try rewrite liftIdRen; try rewrite H; try rewrite H0; try rewrite H1); auto).
|
||||
Qed.
|
||||
|
||||
Lemma idRenDef : forall ty env, idRen (ty :: env) = consMap (ZVAR _ _) (shiftRen ty (idRen env)).
|
||||
intros. apply MapExtensional. intros. dependent destruction var; auto. Qed.
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Composition of renaming
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeRen P env env' env'' (m : Map P env' env'') (r : Ren env env') : Map P env env'' := fun t var => m _ (r _ var).
|
||||
|
||||
Lemma liftComposeRen P ops E env' env'' t (m:Map P env' env'') (r:Ren E env') : lift ops t (composeRen m r) = composeRen (lift ops t m) (liftRen t r).
|
||||
apply MapExtensional. intros t0 var.
|
||||
dependent destruction var; auto.
|
||||
Qed.
|
||||
|
||||
Lemma applyComposeRen env :
|
||||
(forall ty (v : Value env ty) P ops env' env'' (m:Map P env' env'') (s : Ren env env'),
|
||||
mapVal ops (composeRen m s) v = mapVal ops m (renameVal s v))
|
||||
/\ (forall ty (e : Exp env ty) P ops env' env'' (m:Map P env' env'') (s : Ren env env'),
|
||||
mapExp ops (composeRen m s) e = mapExp ops m (renameExp s e)).
|
||||
move: env.
|
||||
apply ExpVal_ind; intros;
|
||||
autorewrite with mapHints renameHints; try rewrite liftComposeRen; try rewrite liftComposeRen; try rewrite <- H;
|
||||
try rewrite <- H0; try rewrite <- H1; auto.
|
||||
Qed.
|
||||
|
||||
Lemma composeRenAssoc :
|
||||
forall P env env' env'' env''' (m : Map P env'' env''') r' (r : Ren env env'),
|
||||
composeRen m (composeRen r' r) = composeRen (composeRen m r') r.
|
||||
by []. Qed.
|
||||
|
||||
Lemma composeRenIdLeft : forall env env' (r : Ren env env'), composeRen (idRen _) r = r.
|
||||
Proof. intros. apply MapExtensional. auto. Qed.
|
||||
|
||||
Lemma composeRenIdRight : forall P env env' (m:Map P env env'), composeRen m (idRen _) = m.
|
||||
Proof. intros. apply MapExtensional. auto. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Substitution
|
||||
==========================================================================*)
|
||||
|
||||
Definition Sub := Map Value.
|
||||
Definition SubMapOps := Build_MapOps TVAR (fun _ _ v => v) (fun env ty ty' => renameVal (fun _ => SVAR ty')).
|
||||
|
||||
(* Convert a renaming into a substitution *)
|
||||
Definition renamingToSub env env' (r : Ren env env') : Sub env env' := fun ty => fun v => TVAR (r ty v).
|
||||
|
||||
Definition subVal := mapVal SubMapOps.
|
||||
Definition subExp := mapExp SubMapOps.
|
||||
|
||||
Definition shiftSub := shiftMap SubMapOps.
|
||||
Implicit Arguments shiftSub [env env'].
|
||||
|
||||
Definition liftSub := lift SubMapOps.
|
||||
Implicit Arguments liftSub [env env'].
|
||||
|
||||
Section SubDefs.
|
||||
|
||||
Variable env env' : Env.
|
||||
Variable s : Sub env env'.
|
||||
Variable t1 t2 : Ty.
|
||||
|
||||
Lemma substTVAR : forall (var : Var env t1), subVal s (TVAR var) = s var. auto. Qed.
|
||||
Lemma substTINT : forall n, subVal s (TINT n) = TINT n. auto. Qed.
|
||||
Lemma substTBOOL : forall n, subVal s (TBOOL n) = TBOOL n. auto. Qed.
|
||||
Lemma substTPAIR : forall (v1 : Value _ t1) (v2:Value _ t2), subVal s (TPAIR v1 v2) = TPAIR (subVal s v1) (subVal s v2). auto. Qed.
|
||||
Lemma substTFST : forall (v : Value _ (t1 ** t2)), subExp s (TFST v) = TFST (subVal s v). auto. Qed.
|
||||
Lemma substTSND : forall (v : Value _ (t1 ** t2)), subExp s (TSND v) = TSND (subVal s v). auto. Qed.
|
||||
Lemma substTFIX : forall (e : Exp (t1 :: t1-->t2 :: _) t2), subVal s (TFIX e) = TFIX (subExp (liftSub _ (liftSub _ s)) e). auto. Qed.
|
||||
Lemma substTOP : forall op v1 v2, subExp s (TOP op v1 v2) = TOP op (subVal s v1) (subVal s v2). auto. Qed.
|
||||
Lemma substTGT : forall v1 v2, subExp s (TGT v1 v2) = TGT (subVal s v1) (subVal s v2). auto. Qed.
|
||||
Lemma substTVAL : forall (v : Value _ t1), subExp s (TVAL v) = TVAL (subVal s v). auto. Qed.
|
||||
Lemma substTLET : forall (e1 : Exp _ t1) (e2 : Exp _ t2), subExp s (TLET e1 e2) = TLET (subExp s e1) (subExp (liftSub _ s) e2). auto. Qed.
|
||||
Lemma substTIF : forall v (e1 e2 : Exp _ t1), subExp s (TIF v e1 e2) = TIF (subVal s v) (subExp s e1) (subExp s e2). auto. Qed.
|
||||
Lemma substTAPP : forall (v1:Value _ (t1 --> t2)) v2, subExp s (TAPP v1 v2) = TAPP (subVal s v1) (subVal s v2). auto. Qed.
|
||||
|
||||
End SubDefs.
|
||||
|
||||
Hint Rewrite substTVAR substTPAIR substTINT substTBOOL substTFST substTSND substTFIX substTOP substTGT substTVAL substTLET substTIF substTAPP : substHints.
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Identity substitution
|
||||
==========================================================================*)
|
||||
|
||||
Definition idSub env : Sub env env := fun ty (x:Var env ty) => TVAR x.
|
||||
Implicit Arguments idSub [].
|
||||
|
||||
Lemma liftIdSub : forall env ty, liftSub ty (idSub env) = idSub (ty :: env).
|
||||
intros env ty. apply MapExtensional. unfold liftSub. intros.
|
||||
dependent destruction var; auto.
|
||||
Qed.
|
||||
|
||||
Lemma applyIdSub env :
|
||||
(forall ty (v : Value env ty), subVal (idSub env) v = v)
|
||||
/\ (forall ty (e : Exp env ty), subExp (idSub env) e = e).
|
||||
Proof.
|
||||
move: env ; apply ExpVal_ind; intros; autorewrite with substHints; try rewrite liftIdSub; try rewrite liftIdSub; try rewrite H; try rewrite H0; try rewrite H1; auto.
|
||||
Qed.
|
||||
|
||||
Notation "[ x , .. , y ]" := (consMap x .. (consMap y (idSub _)) ..) : Sub_scope.
|
||||
Delimit Scope Sub_scope with subst.
|
||||
Arguments Scope subExp [_ _ Sub_scope].
|
||||
Arguments Scope subVal [_ _ Sub_scope].
|
||||
|
||||
Lemma LiftSubDef : forall env env' ty (s : Sub env' env), liftSub ty s = consMap (TVAR (ZVAR _ _)) (shiftSub ty s).
|
||||
intros. apply MapExtensional. intros. dependent destruction var; auto. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Composition of substitution followed by renaming
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeRenSub env env' env'' (r : Ren env' env'') (s : Sub env env') : Sub env env'' := fun t var => renameVal r (s _ var).
|
||||
|
||||
Lemma liftComposeRenSub : forall E env' env'' t (r:Ren env' env'') (s:Sub E env'), liftSub t (composeRenSub r s) = composeRenSub (liftRen t r) (liftSub t s).
|
||||
intros. apply MapExtensional. intros t0 var. dependent induction var; auto.
|
||||
simpl. unfold composeRenSub. unfold liftSub. simpl. unfold renameVal at 1.
|
||||
rewrite <- (proj1 (applyComposeRen _)). unfold renameVal. rewrite <- (proj1 (applyComposeRen _)).
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma applyComposeRenSub E :
|
||||
(forall t (v:Value E t) env' env'' (r:Ren env' env'') (s : Sub E env'),
|
||||
subVal (composeRenSub r s) v = renameVal r (subVal s v))
|
||||
/\ (forall t (e:Exp E t) env' env'' (r:Ren env' env'') (s : Sub E env'),
|
||||
subExp (composeRenSub r s) e = renameExp r (subExp s e)).
|
||||
Proof.
|
||||
move: E ; apply ExpVal_ind;
|
||||
(intros; try (autorewrite with substHints renameHints using try rewrite liftComposeRenSub; try rewrite liftComposeRenSub; try rewrite H; try rewrite H0; try rewrite H1); auto).
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Composition of substitutions
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeSub env env' env'' (s' : Sub env' env'') (s : Sub env env') : Sub env env'' := fun ty var => subVal s' (s _ var).
|
||||
|
||||
Arguments Scope composeSub [_ _ _ Sub_scope Sub_scope].
|
||||
|
||||
Lemma liftComposeSub : forall env env' env'' ty (s' : Sub env' env'') (s : Sub env env'), liftSub ty (composeSub s' s) = composeSub (liftSub ty s') (liftSub ty s).
|
||||
intros. apply MapExtensional. intros t0 var. dependent destruction var. auto.
|
||||
unfold composeSub. simpl.
|
||||
rewrite <- (proj1 (applyComposeRenSub _)). unfold composeRenSub. unfold subVal.
|
||||
rewrite <- (proj1 (applyComposeRen _)). auto.
|
||||
Qed.
|
||||
|
||||
Lemma substComposeSub env :
|
||||
(forall ty (v : Value env ty) env' env'' (s' : Sub env' env'') (s : Sub env env'),
|
||||
subVal (composeSub s' s) v = subVal s' (subVal s v))
|
||||
/\ (forall ty (e : Exp env ty) env' env'' (s' : Sub env' env'') (s : Sub env env'),
|
||||
subExp (composeSub s' s) e = subExp s' (subExp s e)).
|
||||
Proof.
|
||||
move: env ; apply ExpVal_ind;
|
||||
(intros; try (autorewrite with substHints using try rewrite liftComposeSub; try rewrite liftComposeSub; try rewrite H; try rewrite H0; try rewrite H1); auto).
|
||||
Qed.
|
||||
|
||||
Lemma composeCons : forall env env' env'' ty (s':Sub env' env'') (s:Sub env env') (v:Value _ ty),
|
||||
composeSub (consMap v s') (liftSub ty s) = consMap v (composeSub s' s).
|
||||
intros. apply MapExtensional. intros t0 var. dependent destruction var. auto.
|
||||
unfold composeSub. simpl. unfold subVal. rewrite <- (proj1 (applyComposeRen _)). unfold composeRen. simpl.
|
||||
assert ((fun (t0:Ty) (var0:Var env' t0) => s' t0 var0) = s') by (apply MapExtensional; auto).
|
||||
rewrite H. auto.
|
||||
Qed.
|
||||
|
||||
Lemma composeSubIdLeft : forall env env' (s : Sub env env'), composeSub (idSub _) s = s.
|
||||
Proof. intros. apply MapExtensional. intros. unfold composeSub. apply (proj1 (applyIdSub _)). Qed.
|
||||
|
||||
Lemma composeSubIdRight : forall env env' (s:Sub env env'), composeSub s (idSub _) = s.
|
||||
Proof. intros. apply MapExtensional. auto.
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Constructors are injective
|
||||
==========================================================================*)
|
||||
|
||||
Lemma TPAIR_injective :
|
||||
forall env t1 t2 (v1 : Value env t1) (v2 : Value env t2) v3 v4,
|
||||
TPAIR v1 v2 = TPAIR v3 v4 ->
|
||||
v1 = v3 /\ v2 = v4.
|
||||
intros. dependent destruction H. auto.
|
||||
Qed.
|
||||
|
||||
Lemma TFST_injective :
|
||||
forall env t1 t2 (v1 : Value env (t1 ** t2)) v2,
|
||||
TFST v1 = TFST v2 ->
|
||||
v1 = v2.
|
||||
intros. dependent destruction H. auto.
|
||||
Qed.
|
||||
|
||||
Lemma TSND_injective :
|
||||
forall env t1 t2 (v1 : Value env (t1 ** t2)) v2,
|
||||
TSND v1 = TSND v2 ->
|
||||
v1 = v2.
|
||||
intros. dependent destruction H. auto.
|
||||
Qed.
|
||||
|
||||
Lemma TVAL_injective :
|
||||
forall env ty (v1 : Value env ty) v2,
|
||||
TVAL v1 = TVAL v2 ->
|
||||
v1 = v2.
|
||||
intros. dependent destruction H. auto.
|
||||
Qed.
|
||||
|
||||
Lemma TFIX_injective :
|
||||
forall env t1 t2 (e1 : Exp (t1 :: _ :: env) t2) e2,
|
||||
TFIX e1 = TFIX e2 ->
|
||||
e1 = e2.
|
||||
intros. dependent destruction H. auto.
|
||||
Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Closed forms
|
||||
==========================================================================*)
|
||||
Lemma ClosedFunction : forall t1 t2 (v : CValue (t1 --> t2)), exists b, v = TFIX b.
|
||||
unfold CValue.
|
||||
intros t1 t2 v.
|
||||
dependent destruction v.
|
||||
(* TVAR case: impossible *)
|
||||
dependent destruction v.
|
||||
(* TFIX case *)
|
||||
exists e. trivial.
|
||||
Qed.
|
||||
|
||||
Lemma ClosedPair : forall t1 t2 (v : CValue (t1 ** t2)), exists v1, exists v2, v = TPAIR v1 v2.
|
||||
unfold CValue.
|
||||
intros t1 t2 v.
|
||||
dependent destruction v.
|
||||
(* TVAR case: impossible *)
|
||||
dependent destruction v.
|
||||
(* TPAIR case *)
|
||||
exists v1. exists v2. trivial.
|
||||
Qed.
|
||||
|
||||
Lemma ClosedInt : forall (v : CValue Int), exists i, v = TINT i.
|
||||
unfold CValue.
|
||||
intros v. dependent destruction v.
|
||||
exists n. trivial.
|
||||
dependent destruction v.
|
||||
Qed.
|
||||
|
||||
Lemma ClosedBool : forall (v : CValue Bool), exists b, v = TBOOL b.
|
||||
unfold CValue.
|
||||
intros v. dependent destruction v.
|
||||
exists b. trivial.
|
||||
dependent destruction v.
|
||||
Qed.
|
||||
|
||||
Unset Implicit Arguments.
|
76
papers/domains-2012/typedopsem.v
Normal file
76
papers/domains-2012/typedopsem.v
Normal file
|
@ -0,0 +1,76 @@
|
|||
(**********************************************************************************
|
||||
* typedopsem.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
Require Import Coq.Program.Equality.
|
||||
Require Import typedlambda.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
|
||||
(*==========================================================================
|
||||
Evaluation relation
|
||||
==========================================================================*)
|
||||
|
||||
(** printing =>> %\ensuremath{\Downarrow}% *)
|
||||
(** printing n1 %\ensuremath{n_1}% *)
|
||||
(** printing n2 %\ensuremath{n_2}% *)
|
||||
(** printing v1 %\ensuremath{v_1}% *)
|
||||
(** printing v2 %\ensuremath{v_2}% *)
|
||||
(** printing e1 %\ensuremath{e_1}% *)
|
||||
(** printing e2 %\ensuremath{e_2}% *)
|
||||
|
||||
Open Scope Sub_scope.
|
||||
Notation "x '<v=' y" := (x <= y)%N (at level 55).
|
||||
Reserved Notation "e =>> v" (at level 70, no associativity).
|
||||
(*=Ev *)
|
||||
Inductive Ev: forall t, CExp t -> CValue t -> Prop :=
|
||||
| e_Val: forall t (v : CValue t), TVAL v =>> v
|
||||
| e_Op: forall op n1 n2, TOP op (TINT n1) (TINT n2) =>> TINT (op n1 n2)
|
||||
| e_Gt : forall n1 n2, TGT (TINT n1) (TINT n2) =>> TBOOL (n2 <v= n1)
|
||||
| e_Fst : forall t1 t2 (v1 : CValue t1) (v2 : CValue t2), TFST (TPAIR v1 v2) =>> v1
|
||||
| e_Snd : forall t1 t2 (v1 : CValue t1) (v2 : CValue t2), TSND (TPAIR v1 v2) =>> v2
|
||||
| e_App : forall t1 t2 e (v1 : CValue t1) (v2 : CValue t2),
|
||||
subExp [ v1, TFIX e ] e =>> v2 -> TAPP (TFIX e) v1 =>> v2
|
||||
| e_Let : forall t1 t2 e1 e2 (v1 : CValue t1) (v2 : CValue t2),
|
||||
e1 =>> v1 -> subExp [ v1 ] e2 =>> v2 -> TLET e1 e2 =>> v2
|
||||
| e_IfTrue : forall t (e1 e2 : CExp t) v, e1 =>> v -> TIF (TBOOL true) e1 e2 =>> v
|
||||
| e_IfFalse : forall t (e1 e2 : CExp t) v, e2 =>> v -> TIF (TBOOL false) e1 e2 =>> v
|
||||
where "e '=>>' v" := (Ev e v).
|
||||
(*=End *)
|
||||
(*==========================================================================
|
||||
Determinacy of evaluation
|
||||
==========================================================================*)
|
||||
|
||||
Lemma Determinacy: forall ty, forall (e : CExp ty) v1, e =>> v1 -> forall v2, e =>> v2 -> v1 = v2.
|
||||
Proof.
|
||||
intros ty e v1 Ev1.
|
||||
induction Ev1.
|
||||
|
||||
(* e_Val *)
|
||||
intros v2 Ev2. dependent destruction Ev2; trivial.
|
||||
(* e_Op *)
|
||||
intros v2 Ev2. dependent destruction Ev2; trivial.
|
||||
(* e_Gt *)
|
||||
intros v2 Ev2. dependent destruction Ev2; trivial.
|
||||
(* e_Fst *)
|
||||
intros v3 Ev3. dependent destruction Ev3; trivial.
|
||||
(* e_Snd *)
|
||||
intros v3 Ev3. dependent destruction Ev3; trivial.
|
||||
(* e_App *)
|
||||
intros v3 Ev3. dependent destruction Ev3; auto.
|
||||
(* e_Let *)
|
||||
intros v3 Ev3. dependent destruction Ev3. rewrite (IHEv1_1 _ Ev3_1) in IHEv1_2. auto.
|
||||
(* e_IfTrue *)
|
||||
intros v3 Ev3. dependent destruction Ev3; auto.
|
||||
(* e_IfFalse *)
|
||||
intros v3 Ev3. dependent destruction Ev3; auto.
|
||||
Qed.
|
||||
|
||||
Unset Implicit Arguments.
|
65
papers/domains-2012/typedsoundness.v
Normal file
65
papers/domains-2012/typedsoundness.v
Normal file
|
@ -0,0 +1,65 @@
|
|||
(**********************************************************************************
|
||||
* typedsoundness.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
Require Import PredomAll.
|
||||
Require Import typeddensem.
|
||||
Require Import typedopsem.
|
||||
Require Import typedsubst.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
|
||||
(*=Soundness *)
|
||||
Theorem Soundness: forall t (e : CExp t) v, e =>> v -> SemExp e =-= eta << SemVal v.
|
||||
(*=End *)
|
||||
Proof.
|
||||
intros t e v Ev.
|
||||
induction Ev.
|
||||
|
||||
(* e_Val *)
|
||||
simpl. trivial.
|
||||
|
||||
(* e_Op *)
|
||||
simpl. rewrite <- comp_assoc. by apply: fmon_eq_intro.
|
||||
|
||||
(* e_Gt *)
|
||||
simpl. by apply: fmon_eq_intro.
|
||||
|
||||
(* e_Fst *)
|
||||
simpl. rewrite <- comp_assoc. by rewrite prod_fun_fst.
|
||||
|
||||
(* e_Snd *)
|
||||
simpl. rewrite <- comp_assoc. by rewrite prod_fun_snd.
|
||||
|
||||
(* e_App *)
|
||||
rewrite <- IHEv. clear IHEv. clear Ev.
|
||||
|
||||
rewrite <- (proj2 (SemCommutesWithSub _) _ e nil (consMap v1 (consMap (TFIX e) (idSub _)))).
|
||||
simpl. rewrite hdConsMap. rewrite -> (terminal_unique _ Id). simpl.
|
||||
have a:= (@prod_fun_eq_compat cpoProdCatType _ _ _ _ _ (FIXP_com (exp_fun (exp_fun (SemExp e)))) _).
|
||||
specialize (a _ (SemVal v1) (SemVal v1) (tset_refl _)). rewrite a. clear a.
|
||||
rewrite - [FIXP << exp_fun _] comp_idL.
|
||||
rewrite - {1} [exp_fun _] comp_idR.
|
||||
rewrite - prod_map_prod_fun. rewrite comp_assoc. rewrite exp_com.
|
||||
rewrite - {1} [SemVal _] comp_idL. rewrite - prod_map_prod_fun. rewrite comp_assoc. by rewrite exp_com.
|
||||
|
||||
(* e_Let *)
|
||||
simpl. rewrite <- IHEv2. clear IHEv2.
|
||||
rewrite <- (proj2 (SemCommutesWithSub _) _ _ _).
|
||||
simpl. rewrite <- (KLEISLIR_unit (SemExp e2) (SemVal v1) _).
|
||||
rewrite <- IHEv1. clear IHEv1.
|
||||
by rewrite <- (terminal_unique Id _).
|
||||
|
||||
(* e_Iftrue *)
|
||||
simpl. rewrite IHEv. by apply: fmon_eq_intro.
|
||||
|
||||
(* e_Iffalse *)
|
||||
simpl. rewrite IHEv. by apply: fmon_eq_intro.
|
||||
Qed.
|
165
papers/domains-2012/typedsubst.v
Normal file
165
papers/domains-2012/typedsubst.v
Normal file
|
@ -0,0 +1,165 @@
|
|||
(**********************************************************************************
|
||||
* typedsubst.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
Require Import PredomAll.
|
||||
Require Import typedlambda.
|
||||
Require Import typeddensem.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
|
||||
(* Need this for nice pretty-printing *)
|
||||
Unset Printing Implicit Defensive.
|
||||
|
||||
(*==========================================================================
|
||||
Semantics of substitution and renaming
|
||||
==========================================================================*)
|
||||
|
||||
Section SEMMAP.
|
||||
|
||||
Variable P : Env -> Ty -> Type.
|
||||
Variable ops : MapOps P.
|
||||
Variable Sem : forall E t, P E t -> SemEnv E =-> SemTy t.
|
||||
Variable SemVl : forall E t (v : P E t), Sem v =-= SemVal (vl ops v).
|
||||
Variable SemVr : forall E t, Sem (vr ops (ZVAR E t)) =-= pi2.
|
||||
Variable SemWk : forall E t (v:P E t) t', Sem (wk ops t' v) =-= Sem v << pi1.
|
||||
|
||||
(* Apply semantic function pointwise to a renaming or substitution *)
|
||||
(*=SemSub *)
|
||||
Fixpoint SemMap E E' : Map P E' E -> SemEnv E =-> SemEnv E' :=
|
||||
match E' with
|
||||
| nil => fun m => terminal_morph (SemEnv E)
|
||||
| _ :: _ => fun m => <| SemMap (tlMap m) , Sem (hdMap m) |>
|
||||
end.
|
||||
(*=End *)
|
||||
|
||||
Lemma SemShift : forall E E' (m : Map P E E') t, SemMap (shiftMap ops t m) =-= SemMap m << pi1.
|
||||
Proof. elim.
|
||||
- move => E' m t. by apply: terminal_unique.
|
||||
- move => t s IH E' m ty.
|
||||
simpl. destruct (consMapInv m) as [r' [var eq]].
|
||||
subst. simpl. rewrite hdConsMap. rewrite tlConsMap.
|
||||
rewrite -> (IH E' r' ty). rewrite prod_fun_compl. unfold shiftMap, hdMap. by rewrite SemWk.
|
||||
Qed.
|
||||
|
||||
Lemma SemLift : forall E E' (m : Map P E E') t, SemMap (lift ops t m) =-= SemMap m >< Id.
|
||||
Proof. move => E E' m t. simpl. unfold tlMap, hdMap. simpl lift. rewrite SemVr. rewrite SemShift.
|
||||
apply: prod_unique. rewrite prod_fun_fst. by rewrite prod_fun_fst.
|
||||
rewrite prod_fun_snd. rewrite prod_fun_snd. by rewrite comp_idL.
|
||||
Qed.
|
||||
|
||||
(*
|
||||
Lemma SemId : forall E E' (m : Map P E E'), SemMap (fun t (v:Var E t) => vr ops v) =-= Id.
|
||||
Proof. elim.
|
||||
- intros. simpl. by apply: terminal_unique.
|
||||
- intros. simpl. apply:prod_unique. simpl. rewrite <- H. simpl in H. rewrite H. simpl. rewrite H. rewrite tlConsMap hdConsMap. rewrite SemShift. rewrite IH. rewrite comp_idL. rewrite SemVr.
|
||||
apply: prod_unique. rewrite prod_fun_fst. by rewrite comp_idR. rewrite prod_fun_snd. by rewrite comp_idR.
|
||||
Qed.
|
||||
*)
|
||||
Lemma SemCommutesWithMap E:
|
||||
(forall t (v : Value E t) E' (r : Map _ E E'), SemVal v << SemMap r =-= SemVal (mapVal ops r v))
|
||||
/\ (forall t (exp : Exp E t) E' (r : Map _ E E'), SemExp exp << SemMap r =-= SemExp (mapExp ops r exp)).
|
||||
Proof. move: E ; apply ExpVal_ind.
|
||||
|
||||
(* TINT *) by [].
|
||||
(* TBOOL *) intros. simpl. by rewrite <- const_com.
|
||||
|
||||
(* TVAR *)
|
||||
intros E ty var E' r.
|
||||
simpl.
|
||||
induction var.
|
||||
simpl. rewrite prod_fun_snd.
|
||||
destruct (consMapInv r) as [r' [v eq]]. subst.
|
||||
simpl. rewrite hdConsMap. unfold mapVal. unfold travVal. by rewrite SemVl.
|
||||
destruct (consMapInv r) as [r' [v eq]]. subst.
|
||||
simpl.
|
||||
specialize (IHvar r').
|
||||
rewrite <- IHvar. rewrite <- comp_assoc.
|
||||
rewrite tlConsMap. by rewrite prod_fun_fst.
|
||||
|
||||
(* TFIX *)
|
||||
intros E t t1 body IH E' s.
|
||||
specialize (IH _ (lift ops _ (lift ops _ s))).
|
||||
rewrite mapTFIX.
|
||||
simpl SemVal. rewrite <- comp_assoc.
|
||||
do 2 rewrite exp_comp. rewrite <- IH. rewrite SemLift. by rewrite SemLift.
|
||||
|
||||
(* TPAIR *)
|
||||
intros E t1 t2 v1 IH1 v2 IH2 E' s. rewrite mapTPAIR. simpl. rewrite <- IH1. rewrite <- IH2. by rewrite <- prod_fun_compl.
|
||||
|
||||
(* TFST *)
|
||||
intros E t1 t2 v IH E' s. rewrite mapTFST. simpl. rewrite <- IH. by repeat (rewrite comp_assoc).
|
||||
|
||||
(* TSND *)
|
||||
intros E t1 t2 v IH E' s. rewrite mapTSND. simpl. rewrite <- IH. by repeat (rewrite comp_assoc).
|
||||
|
||||
(* TOP *)
|
||||
intros E n v1 IH1 v2 IH2 E' s. rewrite mapTOP. simpl.
|
||||
repeat (rewrite <- comp_assoc). rewrite <- IH1. rewrite <- IH2. by rewrite prod_fun_compl.
|
||||
|
||||
(* TGT *)
|
||||
intros E v1 IH1 v2 IH2 E' s. rewrite mapTGT. simpl.
|
||||
repeat (rewrite <- comp_assoc). rewrite <- IH1. rewrite <- IH2. by rewrite prod_fun_compl.
|
||||
|
||||
(* TVAL *)
|
||||
intros E t v IH E' s. rewrite mapTVAL. simpl.
|
||||
rewrite <- IH. by rewrite <- comp_assoc.
|
||||
|
||||
(* TLET *)
|
||||
intros E t1 t2 e2 IH2 e1 IH1 E' s. rewrite mapTLET. simpl.
|
||||
rewrite <- comp_assoc.
|
||||
specialize (IH2 _ s).
|
||||
specialize (IH1 _ (lift ops _ s)).
|
||||
rewrite prod_fun_compl. rewrite KLEISLIR_comp.
|
||||
rewrite <- IH2. rewrite <- IH1. rewrite SemLift.
|
||||
by do 2 rewrite comp_idL.
|
||||
|
||||
(* TAPP *)
|
||||
intros E t1 t2 v1 IH1 v2 IH2 E' s. rewrite mapTAPP. simpl.
|
||||
rewrite <- comp_assoc. rewrite <- IH1. rewrite <- IH2. by rewrite prod_fun_compl.
|
||||
|
||||
(* TIF *)
|
||||
intros E t ec IHc te1 IH1 te2 IH2 E' s. rewrite mapTIF. simpl.
|
||||
rewrite choose_comp. rewrite -> IH1. rewrite -> IH2. by rewrite -> IHc.
|
||||
Qed.
|
||||
|
||||
End SEMMAP.
|
||||
|
||||
Definition SemRen := SemMap SemVar.
|
||||
Definition SemSub := SemMap SemVal.
|
||||
|
||||
Lemma SemCommutesWithRen E:
|
||||
(forall t (v : Value E t) E' (r : Ren E E'), SemVal v << SemRen r =-= SemVal (renameVal r v))
|
||||
/\ (forall t (e : Exp E t) E' (r : Ren E E'), SemExp e << SemRen r =-= SemExp (renameExp r e)).
|
||||
Proof. by apply SemCommutesWithMap. Qed.
|
||||
|
||||
Lemma SemShiftRen : forall E E' (r : Ren E E') t, SemRen (shiftRen t r) =-= SemRen r << pi1.
|
||||
Proof. by apply SemShift. Qed.
|
||||
|
||||
Lemma SemIdRen : forall E, SemRen (idRen E) =-= Id.
|
||||
elim.
|
||||
- simpl. by apply: terminal_unique.
|
||||
- move => t E IH. simpl. rewrite idRenDef. rewrite tlConsMap.
|
||||
rewrite SemShiftRen. rewrite IH. rewrite comp_idL.
|
||||
by apply: prod_unique ; [rewrite prod_fun_fst | rewrite prod_fun_snd] ; rewrite comp_idR.
|
||||
Qed.
|
||||
|
||||
(*=Substitution *)
|
||||
Lemma SemCommutesWithSub E:
|
||||
(forall t (v : Value E t) E' (s : Sub E E'), SemVal v << SemSub s =-= SemVal (subVal s v))
|
||||
/\ (forall t (e : Exp E t) E' (s : Sub E E'), SemExp e << SemSub s =-= SemExp (subExp s e)).
|
||||
(*=End *)
|
||||
move: E. apply SemCommutesWithMap.
|
||||
+ by []. + by []. + intros. simpl.
|
||||
rewrite <- (proj1 (SemCommutesWithRen E)). assert (SSR := SemShiftRen). unfold shiftRen in SSR. unfold shiftMap in SSR. simpl in SSR.
|
||||
specialize (SSR E E (fun t v => v) t'). simpl in SSR.
|
||||
assert ((fun ty' (var : Var E ty') => SVAR t' var) = (fun t0 => SVAR t')). apply MapExtensional. auto. rewrite H in SSR. rewrite SSR.
|
||||
rewrite SemIdRen. by rewrite comp_idL.
|
||||
Qed.
|
222
papers/domains-2012/unii.v
Executable file
222
papers/domains-2012/unii.v
Executable file
|
@ -0,0 +1,222 @@
|
|||
(**********************************************************************************
|
||||
* unii.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Unityped lambda calculus, well-scoped by construction *)
|
||||
|
||||
Require Export ssreflect ssrnat.
|
||||
Require Import Program.
|
||||
Require Import Fin.
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
Ltac Rewrites E :=
|
||||
(intros; simpl; try rewrite E;
|
||||
repeat (match goal with | [H : context[eq _ _] |- _] => rewrite H end);
|
||||
auto).
|
||||
|
||||
Definition Env := nat.
|
||||
|
||||
Inductive Value E :=
|
||||
| VAR: Fin E -> Value E
|
||||
| INT: nat -> Value E
|
||||
| LAMBDA: Exp E.+1 -> Value E
|
||||
with Exp E :=
|
||||
| VAL: Value E -> Exp E
|
||||
| APP: Value E -> Value E -> Exp E
|
||||
| LET: Exp E -> Exp E.+1 -> Exp E
|
||||
| IFZ: Value E -> Exp E -> Exp E -> Exp E
|
||||
| OP: (nat -> nat -> nat) -> Value E -> Value E -> Exp E.
|
||||
|
||||
Implicit Arguments INT [E].
|
||||
|
||||
Scheme Value_ind2 := Induction for Value Sort Prop
|
||||
with Exp_ind2 := Induction for Exp Sort Prop.
|
||||
Combined Scheme ExpValue_ind from Value_ind2, Exp_ind2.
|
||||
|
||||
(*==========================================================================
|
||||
Variable-domain maps.
|
||||
By instantiating P with Var we get renamings.
|
||||
By instantiating P with Value we get substitutions.
|
||||
==========================================================================*)
|
||||
Module Map.
|
||||
Section MAP.
|
||||
|
||||
Variable P : Env -> Type.
|
||||
Definition Map E E' := FMap E (P E').
|
||||
|
||||
(*==========================================================================
|
||||
Package of operations used with a Map
|
||||
vr maps a Var into Var or Value (so is either the identity or TVAR)
|
||||
vl maps a Var or Value to a Value (so is either TVAR or the identity)
|
||||
wk weakens a Var or Value (so is either SVAR or renaming through SVAR on a value)
|
||||
==========================================================================*)
|
||||
Record Ops :=
|
||||
{
|
||||
vr : forall E, Fin E -> P E;
|
||||
vl : forall E, P E -> Value E;
|
||||
wk : forall E, P E -> P E.+1;
|
||||
wkvr : forall E (var : Fin E), wk (vr var) = vr (FinS var);
|
||||
vlvr : forall E (var : Fin E), vl (vr var) = VAR var
|
||||
}.
|
||||
Variable ops : Ops.
|
||||
|
||||
Definition shift E E' (m : Map E E') : Map E E'.+1 := fun var => wk ops (m var).
|
||||
Definition id E : Map E E := fun (var : Fin E) => vr ops var.
|
||||
Definition lift E E' (m : Map E E') : Map E.+1 E'.+1 := cons (vr ops (FinZ _)) (shift m).
|
||||
|
||||
Lemma shiftCons : forall E E' (m : Map E E') (x : P E'), shift (cons x m) = cons (wk ops x) (shift m).
|
||||
Proof. move => E E' m x. extFMap var; by []. Qed.
|
||||
|
||||
Fixpoint mapVal E E' (m : Map E E') (v : Value E) : Value E' :=
|
||||
match v with
|
||||
| VAR v => vl ops (m v)
|
||||
| INT i => INT i
|
||||
| LAMBDA e => LAMBDA (mapExp (lift m) e)
|
||||
end
|
||||
with mapExp E E' (m : Map E E') (e : Exp E) : Exp E' :=
|
||||
match e with
|
||||
| VAL v => VAL (mapVal m v)
|
||||
| APP v0 v1 => APP (mapVal m v0) (mapVal m v1)
|
||||
| LET e0 e1 => LET (mapExp m e0) (mapExp (lift m) e1)
|
||||
| OP op v0 v1 => OP op (mapVal m v0) (mapVal m v1)
|
||||
| IFZ v e0 e1 => IFZ (mapVal m v) (mapExp m e0) (mapExp m e1)
|
||||
end.
|
||||
|
||||
Variable E E' : Env.
|
||||
Variable m : Map E E'.
|
||||
Lemma mapVAR : forall (var : Fin _), mapVal m (VAR var) = vl ops (m var). by []. Qed.
|
||||
Lemma mapINT : forall n, mapVal m (INT n) = INT n. by []. Qed.
|
||||
Lemma mapLAMBDA : forall (e : Exp _), mapVal m (LAMBDA e) = LAMBDA (mapExp (lift m) e). by []. Qed.
|
||||
Lemma mapOP : forall op v1 v2, mapExp m (OP op v1 v2) = OP op (mapVal m v1) (mapVal m v2). by []. Qed.
|
||||
Lemma mapVAL : forall (v : Value _), mapExp m (VAL v) = VAL (mapVal m v). by []. Qed.
|
||||
Lemma mapLET : forall (e1 : Exp _) (e2 : Exp _), mapExp m (LET e1 e2) = LET (mapExp m e1) (mapExp (lift m) e2). by []. Qed.
|
||||
Lemma mapIFZ : forall v (e1 e2 : Exp _), mapExp m (IFZ v e1 e2) = IFZ (mapVal m v) (mapExp m e1) (mapExp m e2). by []. Qed.
|
||||
Lemma mapAPP : forall (v1 : Value _) v2, mapExp m (APP v1 v2) = APP (mapVal m v1) (mapVal m v2). by []. Qed.
|
||||
|
||||
Lemma liftId : lift (@id E) = @id E.+1.
|
||||
Proof. extFMap var; [by [] | by apply wkvr].
|
||||
Qed.
|
||||
|
||||
Lemma idAsCons : @id E.+1 = cons (vr ops (FinZ _)) (shift (@id E)).
|
||||
Proof. extFMap var; first by []. unfold id, shift. simpl. by rewrite wkvr. Qed.
|
||||
|
||||
End MAP.
|
||||
|
||||
Hint Rewrite mapVAR mapINT mapLAMBDA mapOP mapVAL mapLET mapIFZ mapAPP : mapHints.
|
||||
Implicit Arguments id [P].
|
||||
|
||||
Lemma applyId P (ops:Ops P) E :
|
||||
(forall (v : Value E), mapVal ops (id ops E) v = v)
|
||||
/\ (forall (e : Exp E), mapExp ops (id ops E) e = e).
|
||||
Proof. move: E ; apply ExpValue_ind; intros; autorewrite with mapHints; Rewrites liftId. by apply vlvr. Qed.
|
||||
|
||||
End Map.
|
||||
|
||||
(*==========================================================================
|
||||
Variable renamings: Map Var
|
||||
==========================================================================*)
|
||||
|
||||
Definition Ren := Map.Map Fin.
|
||||
Definition RenOps : Map.Ops Fin. refine (@Map.Build_Ops _ (fun _ v => v) VAR FinS _ _). by []. by []. Defined.
|
||||
|
||||
Definition renVal := Map.mapVal RenOps.
|
||||
Definition renExp := Map.mapExp RenOps.
|
||||
Definition liftRen := Map.lift RenOps.
|
||||
Definition shiftRen := Map.shift RenOps.
|
||||
Definition idRen := Map.id RenOps.
|
||||
|
||||
(*==========================================================================
|
||||
Composition of renaming
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeRen P E E' E'' (m : Map.Map P E' E'') (r : Ren E E') : Map.Map P E E'' := fun var => m (r var).
|
||||
|
||||
Lemma liftComposeRen : forall E E' E'' P ops (m:Map.Map P E' E'') (r:Ren E E'), Map.lift ops (composeRen m r) = composeRen (Map.lift ops m) (liftRen r).
|
||||
Proof. move => E E' E'' P ops m r. extFMap var; by []. Qed.
|
||||
|
||||
Lemma applyComposeRen E :
|
||||
(forall (v : Value E) E' E'' P ops (m:Map.Map P E' E'') (s : Ren E E'),
|
||||
Map.mapVal ops (composeRen m s) v = Map.mapVal ops m (renVal s v))
|
||||
/\ (forall (e : Exp E) E' E'' P ops (m:Map.Map P E' E'') (s : Ren E E'),
|
||||
Map.mapExp ops (composeRen m s) e = Map.mapExp ops m (renExp s e)).
|
||||
Proof. move: E ; apply ExpValue_ind; intros; autorewrite with mapHints; Rewrites liftComposeRen. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Substitution
|
||||
==========================================================================*)
|
||||
|
||||
Definition Sub := Map.Map Value.
|
||||
Definition SubOps : Map.Ops Value. refine (@Map.Build_Ops _ VAR (fun _ v => v) (fun E => renVal (fun v => FinS v)) _ _). by []. by []. Defined.
|
||||
|
||||
Definition subVal := Map.mapVal SubOps.
|
||||
Definition subExp := Map.mapExp SubOps.
|
||||
Definition shiftSub := Map.shift SubOps.
|
||||
Definition liftSub := Map.lift SubOps.
|
||||
Definition idSub := Map.id SubOps.
|
||||
|
||||
Ltac UnfoldRenSub := (unfold subVal; unfold subExp; unfold renVal; unfold renExp; unfold liftSub; unfold liftRen).
|
||||
Ltac FoldRenSub := (fold subVal; fold subExp; fold renVal; fold renExp; fold liftSub; fold liftRen).
|
||||
Ltac SimplMap := (UnfoldRenSub; autorewrite with mapHints; FoldRenSub).
|
||||
|
||||
(*==========================================================================
|
||||
Composition of substitution followed by renaming
|
||||
==========================================================================*)
|
||||
Definition composeRenSub E E' E'' (r : Ren E' E'') (s : Sub E E') : Sub E E'' := fun var => renVal r (s var).
|
||||
|
||||
Lemma liftComposeRenSub : forall E E' E'' (r:Ren E' E'') (s:Sub E E'), liftSub (composeRenSub r s) = composeRenSub (liftRen r) (liftSub s).
|
||||
intros. extFMap var; first by [].
|
||||
unfold composeRenSub, liftSub. simpl. unfold renVal at 1 3. by do 2 rewrite <- (proj1 (applyComposeRen _)).
|
||||
Qed.
|
||||
|
||||
Lemma applyComposeRenSub E :
|
||||
(forall (v:Value E) E' E'' (r:Ren E' E'') (s : Sub E E'),
|
||||
subVal (composeRenSub r s) v = renVal r (subVal s v))
|
||||
/\ (forall (e:Exp E) E' E'' (r:Ren E' E'') (s : Sub E E'),
|
||||
subExp (composeRenSub r s) e = renExp r (subExp s e)).
|
||||
Proof. move: E ; apply ExpValue_ind; intros; SimplMap; Rewrites liftComposeRenSub. Qed.
|
||||
|
||||
(*==========================================================================
|
||||
Composition of substitutions
|
||||
==========================================================================*)
|
||||
|
||||
Definition composeSub E E' E'' (s' : Sub E' E'') (s : Sub E E') : Sub E E'' := fun var => subVal s' (s var).
|
||||
|
||||
Lemma liftComposeSub : forall E E' E'' (s' : Sub E' E'') (s : Sub E E'), liftSub (composeSub s' s) = composeSub (liftSub s') (liftSub s).
|
||||
intros. extFMap var; first by [].
|
||||
unfold composeSub. simpl. rewrite <- (proj1 (applyComposeRenSub _)). unfold composeRenSub, subVal. by rewrite <- (proj1 (applyComposeRen _)).
|
||||
Qed.
|
||||
|
||||
Lemma applyComposeSub E :
|
||||
(forall (v : Value E) E' E'' (s' : Sub E' E'') (s : Sub E E'),
|
||||
subVal (composeSub s' s) v = subVal s' (subVal s v))
|
||||
/\ (forall (e : Exp E) E' E'' (s' : Sub E' E'') (s : Sub E E'),
|
||||
subExp (composeSub s' s) e = subExp s' (subExp s e)).
|
||||
Proof. move: E ; apply ExpValue_ind; intros; SimplMap; Rewrites liftComposeSub. Qed.
|
||||
|
||||
Lemma composeCons : forall E E' E'' (s':Sub E' E'') (s:Sub E E') (v:Value _),
|
||||
composeSub (cons v s') (liftSub s) = cons v (composeSub s' s).
|
||||
Proof. intros. extFMap var; first by [].
|
||||
unfold composeSub. unfold subVal. simpl. rewrite <- (proj1 (applyComposeRen _)). unfold composeRen.
|
||||
simpl. replace ((fun (var0:Fin E') => s' var0)) with s' by (by apply Extensionality). by [].
|
||||
Qed.
|
||||
|
||||
Lemma composeSubIdLeft : forall E E' (s : Sub E E'), composeSub (@idSub _) s = s.
|
||||
Proof. intros. extFMap var; by apply (proj1 (Map.applyId _ _)). Qed.
|
||||
|
||||
Lemma composeSubIdRight : forall E E' (s:Sub E E'), composeSub s (@idSub _) = s.
|
||||
Proof. intros. extFMap var; by []. Qed.
|
||||
|
||||
Notation "[ x , .. , y ]" := (cons x .. (cons y (@Map.id _ SubOps _)) ..) : Sub_scope.
|
||||
Arguments Scope composeSub [_ _ _ Sub_scope Sub_scope].
|
||||
Arguments Scope subExp [_ _ Sub_scope].
|
||||
Arguments Scope subVal [_ _ Sub_scope].
|
||||
Delimit Scope Sub_scope with sub.
|
||||
|
||||
Lemma composeSingleSub : forall E E' (s:Sub E E') (v:Value _), composeSub [v] (liftSub s) = cons v s.
|
||||
Proof. intros. rewrite composeCons. by rewrite composeSubIdLeft. Qed.
|
543
papers/domains-2012/uniiade.v
Executable file
543
papers/domains-2012/uniiade.v
Executable file
|
@ -0,0 +1,543 @@
|
|||
(**********************************************************************************
|
||||
* uniiade.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Adequacy of semantics for unityped lambda calculus *)
|
||||
|
||||
Require Import unii.
|
||||
Require Export uniisem uniiop Fin.
|
||||
Require Import PredomAll.
|
||||
Require Import KnasterTarski.
|
||||
Require Import NSetoid.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
Open Scope C_scope.
|
||||
|
||||
Definition RelKind : setoidType := CPO.setoidType VInf * discrete_setoidType (Value O).
|
||||
|
||||
Lemma RelV_respect (R S:RelKind =-> Props) : setoid_respect
|
||||
(fun (p:RelKind) => match p with (d,v) => match v with
|
||||
| INT m => d =-= inl m
|
||||
| VAR _ => False
|
||||
| LAMBDA e =>
|
||||
exists f, d =-= inr f /\
|
||||
forall (d1:VInf) (v1:Value O), R (d1,v1) ->
|
||||
forall dv, kleisli (eta << Unroll) (f (Roll d1)) =-= Val dv ->
|
||||
exists v2, subExp ([v1]) e =>> v2 /\ S (dv,v2)
|
||||
end end).
|
||||
case => d1 v1. case => d2 v2. case => e0 e1. simpl in e0,e1. case: e1 => e1.
|
||||
rewrite e1 ; clear e1 v1. case: v2 ; first by [].
|
||||
- move => n. split => e ; by rewrite <- e ; auto.
|
||||
- move => e. by split ; case => f [e1 P] ; exists f ; rewrite <- e1 ; split ; auto.
|
||||
Qed.
|
||||
|
||||
Definition RelV (R S:RelKind =-> Props) : (RelKind =-> Props) :=
|
||||
Eval hnf in mk_fset (RelV_respect R S).
|
||||
|
||||
Definition imageD (f:VInf =-> VInf _BOT) (R S:RelKind =-> Props) :=
|
||||
(forall x, R x -> let (d,v) := x in forall dd, (f d) =-= Val dd -> S (dd,v)).
|
||||
|
||||
Definition RAdm (R : RelKind =-> Props) := forall (S : RelKind =-> Props)
|
||||
(c : natO =-> (VInf -=> VInf _BOT)), (forall n, imageD (c n) S R) ->
|
||||
imageD (lub c) S R.
|
||||
|
||||
Lemma RAdm_inf : forall S : set_clatType RelKind -> Prop,
|
||||
(forall t : set_clatType RelKind, S t -> RAdm t) -> RAdm (@inf (set_clatType RelKind) S).
|
||||
intros S C. unfold RAdm in *.
|
||||
intros R c A. unfold imageD. intros de. case de ; clear de. intros d e Rd dd ldd.
|
||||
simpl. intros Si SSi.
|
||||
specialize (C _ SSi R c).
|
||||
assert ((forall n : natO, imageD (c n) R Si)) as xx.
|
||||
intros nn. specialize (A nn). simpl in A. unfold imageD in A. unfold imageD.
|
||||
intros x ; case x ; clear x. intros d1 v1 R1. specialize (A _ R1). simpl in A.
|
||||
intros dd1 ddv1. specialize (A _ ddv1 _ SSi). apply A.
|
||||
specialize (C xx). unfold imageD in C.
|
||||
specialize (C _ Rd _ ldd). apply C.
|
||||
Qed.
|
||||
|
||||
Definition RelAdm := @sub_clatType (set_clatType RelKind) RAdm RAdm_inf.
|
||||
|
||||
Definition RelAdmop := op_latType RelAdm.
|
||||
|
||||
Lemma ev_determ: forall e v1, e =>> v1 -> forall v2, e =>> v2 -> v1 = v2.
|
||||
intros v v1 ev. induction ev ; try (intros ; inversion H ; auto ; fail).
|
||||
intros v3 ev. inversion ev. subst. rewrite <- (IHev1 _ H1) in *. apply (IHev2 _ H3).
|
||||
Qed.
|
||||
|
||||
Lemma sum_right_simpl: forall (D E:cpoType) (c:natO =-> (D + E)) T k e, c k =-= inr e -> sum_right (D1:=D) (c:=c) T k =-= e.
|
||||
intros D E c T k e. case T. simpl. clear T. intros x c0. clear c0. case: (c k). move => s. by case.
|
||||
move => s A. by apply A.
|
||||
Qed.
|
||||
|
||||
Lemma single_set_respect (T:setoidType) (v:T) : setoid_respect (fun y => y =-= v).
|
||||
move => y y' e.
|
||||
by split => e' ; rewrite <- e' ; rewrite e.
|
||||
Qed.
|
||||
|
||||
Definition single_set (T:setoidType) (v:T) : T =-> Props := Eval hnf in mk_fset (single_set_respect v).
|
||||
|
||||
Lemma Kab_mono (D1 D2:cpoType) : monotonic (fun x:D2 => @const D1 D2 x:D1 -=> D2).
|
||||
by move => x y l.
|
||||
Qed.
|
||||
|
||||
Definition Kab (D1 D2:cpoType) : ordCatType D2 (D1 -=> D2) := Eval hnf in mk_fmono (@Kab_mono D1 D2).
|
||||
|
||||
Lemma inv_image_respect (T T':setoidType) (f: T =-> T') (Y:T' =-> Props) : setoid_respect (fun x => Y (f x)).
|
||||
move => x y e. simpl. by rewrite -> e.
|
||||
Qed.
|
||||
|
||||
Definition inv_image (T T':setoidType) (f: T =-> T') (Y:T' =-> Props) : (T =-> Props) :=
|
||||
Eval hnf in mk_fset (inv_image_respect f Y).
|
||||
|
||||
Lemma RelVA_adm (R : set_clatType RelKind)
|
||||
(AR : RAdm R) (S : set_clatType RelKind) (AS : RAdm S) : RAdm (RelV R S).
|
||||
unfold RAdm.
|
||||
simpl. move => A c C.
|
||||
unfold imageD. intros e ; case e ; clear e. intros d v Ad dd lubdd.
|
||||
simpl in lubdd.
|
||||
assert (xx:=lubval lubdd). destruct xx as [k [dd' [ck Pdd]]].
|
||||
destruct (DLvalgetchain ck) as [c' Pc'].
|
||||
|
||||
generalize Ad. clear Ad.
|
||||
case v ; clear v. simpl. intros n. Require Import Program. by dependent destruction n.
|
||||
|
||||
intros n cA. simpl. specialize (C k). unfold imageD in C.
|
||||
specialize (C _ cA _ ck).
|
||||
assert (lub (fcont_app c d) =-= Val (lub c')) as lc.
|
||||
rewrite <- eta_val. rewrite -> (lub_comp_eq _ c').
|
||||
rewrite -> (lub_lift_left _ k).
|
||||
refine (lub_eq_compat _). refine (fmon_eq_intro _) => m.
|
||||
simpl. apply (Pc' m). simpl in lc. rewrite -> lc in lubdd.
|
||||
assert (ll := vinj lubdd). clear lubdd.
|
||||
rewrite <- ll in *. clear ll. simpl in C. rewrite -> C in Pdd.
|
||||
assert (c' 0%N =-= @INL (discrete_cpoType nat) (DInf -=> DInf _BOT) n).
|
||||
specialize (Pc' O). rewrite addn0 in Pc'.
|
||||
rewrite -> ck in Pc'. assert (xx := vinj Pc'). rewrite <- xx. rewrite -> C. by auto.
|
||||
assert (forall x, @INL (discrete_cpoType nat) (DInf -=> DInf _BOT) n <= x -> exists d, x =-= inl d).
|
||||
intros x. simpl. case x. intros t nt. simpl. exists t ; auto.
|
||||
intros t FF. inversion FF. specialize (H0 _ Pdd).
|
||||
destruct H0 as [nn sn].
|
||||
rewrite -> sn in Pdd. simpl in Pdd. rewrite <- Pdd in *. by apply sn.
|
||||
|
||||
simpl.
|
||||
intros ee Cc. assert (C0 := C k). unfold imageD in C0. specialize (C0 _ Cc). simpl in C0.
|
||||
specialize (C0 _ ck).
|
||||
|
||||
destruct (C0) as [f [cf _]].
|
||||
case (SuminjProof c' 0). intros incon. case incon. clear incon.
|
||||
intros ff incon. specialize (Pc' O). rewrite -> addn0 in Pc'. rewrite incon in Pc'. simpl in Pc'.
|
||||
rewrite -> ck in Pc'. have Pcc:=vinj Pc'. rewrite -> cf in Pcc. by case: Pcc.
|
||||
|
||||
intros T. exists (lub (sum_right T)).
|
||||
assert (dd =-=
|
||||
inr (* nat *)
|
||||
(@lub (DInf -=> DInf _BOT)
|
||||
(sum_right (D1:=discrete_cpoType nat) (D2:=DInf -=> DInf _BOT) (c:=c') T))) as ldd.
|
||||
have CI:(continuous ( (@INR (discrete_cpoType nat) (DInf -=> DInf _BOT)))) by auto.
|
||||
assert (forall x, Inr (discrete_ordType nat) (DInf -=> DInf _BOT) x = inr x).
|
||||
intros x. by auto. rewrite <- H. clear H. rewrite -> (lub_comp_eq (INR (discrete_cpoType nat) _) (sum_right T)).
|
||||
|
||||
assert (lub (fcont_app c d) =-= Val (lub c')) as lc.
|
||||
rewrite <- eta_val. rewrite -> (lub_comp_eq _ c').
|
||||
refine (Oeq_trans (lub_lift_left (fcont_app c d) k) _).
|
||||
refine (lub_eq_compat _). refine (fmon_eq_intro _).
|
||||
intros m. simpl. apply (Pc' m). simpl in lc. rewrite -> lc in lubdd.
|
||||
assert (ll := vinj lubdd). clear lubdd.
|
||||
rewrite <- ll. assert (xx:=@lub_eq_compat _ c'). simpl in xx. simpl. apply xx.
|
||||
refine (fmon_eq_intro _). intros m. simpl.
|
||||
assert (exists dd, c' m =-= inr (* (discrete_cpoType nat) *) dd). assert (c' O <= c' m) as cm.
|
||||
assert (monotonic c') as cm' by auto. by apply cm'. simpl. case T. intros x c'0.
|
||||
simpl in c'0.
|
||||
generalize cm. clear cm. case (c' m). intros t incon. simpl in incon. rewrite c'0 in incon. inversion incon.
|
||||
intros t cm. exists t. auto. destruct H as [c'm cm].
|
||||
assert (yy := sum_right_simpl T cm). rewrite -> cm. simpl.
|
||||
assert (stable (@INR (discrete_cpoType nat) (DInf -=> DInf _BOT))) by auto.
|
||||
simpl in H. refine (H _ _ _). clear H.
|
||||
rewrite <- yy. by auto.
|
||||
|
||||
split. apply ldd.
|
||||
|
||||
intros d1 v1 Sd1 dv rdv.
|
||||
rewrite fcont_lub_simpl in rdv.
|
||||
rewrite -> (lub_comp_eq _ (fcont_app (sum_right T) (Roll d1))) in rdv.
|
||||
destruct (lubval rdv) as [kk [dkk [ckk dkkl]]].
|
||||
destruct (DLvalgetchain ckk) as [c1' cP]. simpl in ckk.
|
||||
|
||||
assert (D := C (k+kk)%N). unfold imageD in D. specialize (D _ Cc). simpl in D.
|
||||
assert (xx:= Pc' kk). specialize (D _ xx).
|
||||
case: D => ff. case => c0 P.
|
||||
specialize (P _ v1 Sd1).
|
||||
assert (cp0:= cP O).
|
||||
|
||||
assert (bla:=sum_right_simpl T c0). simpl in cp0. rewrite addn0 in cp0.
|
||||
assert (stable ((kleisli (eta << Unroll)))) as su by apply: fmon_stable.
|
||||
specialize (su _ _ (fmon_eq_elim bla (Roll (d1)))). rewrite -> su in cp0.
|
||||
specialize (P _ cp0). destruct P as [v2 [ev2 S2]].
|
||||
exists v2. split. by apply ev2. clear bla su cp0 xx S2 c0 ff.
|
||||
|
||||
unfold RAdm in AS.
|
||||
|
||||
specialize (AS (inv_image (@Ssnd _ _) (single_set (v2:discrete_setoidType (Value O))))).
|
||||
specialize (AS (Kab _ (VInf _BOT) << (eta_m << c1'))).
|
||||
have AA:(forall n : natO,
|
||||
imageD ((Kab VInf (VInf _BOT) << (eta_m << c1')) n)
|
||||
(inv_image
|
||||
(Ssnd (CPO.setoidType VInf) (discrete_setoidType (Value 0)))
|
||||
(single_set (v2:discrete_setoidType (Value 0)))) S).
|
||||
intros n. unfold imageD. simpl. intros x ; case x; clear x. simpl.
|
||||
intros d0 v0 veq. unfold tset_eq in veq. simpl in veq. rewrite veq. clear v0 veq. intros d00.
|
||||
move => vv. have vvv:=vinj vv. clear vv.
|
||||
specialize (C (k+kk+n)%N).
|
||||
unfold imageD in C. specialize (C _ Cc). simpl in C.
|
||||
specialize (cP n).
|
||||
assert (c (k + kk + n)%N d =-= Val (c' (kk + n))%N).
|
||||
specialize (Pc' (kk+n)%N). rewrite <- Pc'.
|
||||
by rewrite addnA.
|
||||
specialize (C _ H).
|
||||
destruct C as [fn [c'n Pn']].
|
||||
specialize (Pn' _ _ Sd1). simpl in cP. clear H.
|
||||
assert (xx:=sum_right_simpl T c'n).
|
||||
assert (stable ( (kleisli (eta << Unroll)))) as su by apply:fmon_stable.
|
||||
specialize (su _ _ (fmon_eq_elim xx (Roll d1))). rewrite -> su in cP.
|
||||
clear su xx. specialize (Pn' _ cP).
|
||||
destruct Pn' as [v3 [ev3 S3]]. have e0:=ev_determ ev3 ev2. rewrite -> e0 in S3. clear ev3 v3 e0.
|
||||
have e0:S (d00, v2) =-= S (c1' n,v2). apply frespect. split ; last by []. simpl. by rewrite -> vvv.
|
||||
apply (proj2 e0). by apply S3. specialize (AS AA).
|
||||
unfold imageD in AS.
|
||||
specialize (AS (d,v2)). simpl in AS.
|
||||
refine (AS _ _ _) ; first by [].
|
||||
assert (xx := fcont_lub_simpl (Kab VInf (VInf _BOT) << (eta_m << c1')) d).
|
||||
simpl in xx. rewrite xx. clear xx.
|
||||
rewrite <- rdv.
|
||||
assert (aa:=@lub_eq_compat (VInf _BOT)).
|
||||
simpl in aa.
|
||||
refine (Oeq_trans _ (Oeq_sym (lub_lift_left _ kk))).
|
||||
refine (aa _ _ _). refine (fmon_eq_intro _). intros n.
|
||||
specialize (cP n). auto using cP.
|
||||
Qed.
|
||||
|
||||
Definition RelVA : RelAdmop -> RelAdm -> RelAdm.
|
||||
intros R S. case R. clear R. intros R AR. case S. clear S. intros S AS.
|
||||
exists (RelV R S). by apply RelVA_adm.
|
||||
Defined.
|
||||
|
||||
Definition RelVAop : RelAdm -> RelAdmop -> RelAdmop := fun X Y => RelVA X Y.
|
||||
|
||||
Lemma RelVm_mon : monotonic (fun (p:prod_clatType RelAdmop RelAdm) =>
|
||||
(RelVAop (lsnd _ _ p) (lfst _ _ p), RelVA (lfst _ _ p) (lsnd _ _ p))).
|
||||
unfold monotonic. intros x. case x. clear x. intros x0 y0 y. case y. clear y. intros x1 y1.
|
||||
simpl. case x1. clear x1. intros x1 rx1. case x0. clear x0. intros x0 rx0.
|
||||
case y0 ; clear y0 ; intros y0 ry0. case y1 ; clear y1 ; intros y1 ry1.
|
||||
intros [sx sy]. split. simpl.
|
||||
intros dv. case dv. clear dv. intros d v. simpl. case v ; [by auto | by auto | idtac].
|
||||
clear v. intros e C. destruct C as [f [df P]]. exists f. split ; first by auto.
|
||||
intros d1 v1 y01. specialize (P d1 v1).
|
||||
specialize (sy (d1,v1) y01). intros dv rdv. specialize (P sy dv rdv).
|
||||
destruct P as [v2 [ev2 P]]. exists v2. split; [by apply ev2 | by apply (sx (dv,v2) P)].
|
||||
|
||||
simpl.
|
||||
intros dv. case dv. clear dv. intros d v. simpl.
|
||||
case v ; auto. clear v. intros v C. destruct C as [f [df C]].
|
||||
exists f. split ; first by auto. intros d1 v1 c1. specialize (C d1 v1).
|
||||
specialize (sx (d1,v1) c1). intros dv dvr. specialize (C sx dv dvr).
|
||||
destruct C as [v2 [ev2 C]]. exists v2. split; [by apply ev2 | by apply (sy (dv,v2) C)].
|
||||
Qed.
|
||||
|
||||
Definition RelVm : ordCatType (prod_clatType RelAdmop RelAdm) (prod_clatType RelAdmop RelAdm) :=
|
||||
Eval hnf in mk_fmono RelVm_mon.
|
||||
|
||||
Definition Delta := lsnd _ _ (lfp RelVm).
|
||||
Definition Deltam := lfst _ _ (lfp RelVm).
|
||||
|
||||
Lemma RelV_fixedD: Delta =-= (RelVA Deltam Delta).
|
||||
unfold Delta, Deltam. assert (xx:=lfp_fixedpoint RelVm).
|
||||
generalize xx. clear xx.
|
||||
case ((@lfp (prod_clatType RelAdmop RelAdm) RelVm)).
|
||||
intros d0 d1. simpl. intros X. case: X. simpl. case => A B. case => X Y. by split.
|
||||
Qed.
|
||||
|
||||
Lemma RelV_fixedDm : Deltam =-= (RelVA Delta Deltam).
|
||||
unfold Deltam,Delta. assert (xx := lfp_fixedpoint RelVm).
|
||||
generalize xx ;clear xx.
|
||||
case (lfp RelVm).
|
||||
intros d0 d1. simpl. case. simpl. case => A B. case => X Y. by split.
|
||||
Qed.
|
||||
|
||||
Lemma RelV_least: forall R S, R <= RelVA S R -> (RelVA R S) <= S ->
|
||||
R <= Deltam /\ Delta <= S.
|
||||
intros R S.
|
||||
assert (xx:=@lfp_least _ RelVm).
|
||||
intros Rsub Ssub. simpl in Rsub, Ssub.
|
||||
assert (RelVm (R,S) <= (R,S)). simpl. split. by apply Rsub. by apply Ssub.
|
||||
specialize (xx _ H).
|
||||
assert (lfp RelVm =-= (Deltam,Delta)) as L.
|
||||
unfold Deltam, Delta. by split.
|
||||
rewrite -> L in xx.
|
||||
by destruct xx ; split ; auto.
|
||||
Qed.
|
||||
|
||||
Require Import uniirec.
|
||||
Lemma RelV_F_action: forall (f:VInf =-> VInf _BOT) (R S:RelKind =-> Props), imageD f R S ->
|
||||
imageD (eta << [|in1,
|
||||
(in2 <<
|
||||
((exp_fun (CCOMP _ _ _:cpoCatType _ _) (kleisli (kleisli (eta << Roll) << f << Unroll)) : cpoCatType _ _) <<
|
||||
((exp_fun
|
||||
((CCOMP DInf _ (DInf _BOT)) <<
|
||||
SWAP) ((kleisli (eta << Roll) << f << Unroll)) : cpoCatType _ _) << KLEISLI)))|]) (RelV S R) (RelV R S).
|
||||
move => f R S X.
|
||||
case => d. case ; first by [].
|
||||
- move => n. simpl. move => e dd. move => A. have A':=vinj A. clear A. rewrite -> e in A'. clear d e.
|
||||
rewrite -> SUM_fun_simplx in A'. by apply (Oeq_sym A').
|
||||
- move => e. simpl. case => f' [e' P] dd ee.
|
||||
exists (kleisli ((kleisli (eta << Roll) << f) << Unroll) <<
|
||||
(kleisli f' << ((kleisli (eta << Roll) << f) << Unroll))).
|
||||
have ee':=vinj ee. clear ee. rewrite -> e' in ee'. rewrite -> SUM_fun_simplx in ee'.
|
||||
simpl in ee'. rewrite <- ee'. split ; first by [].
|
||||
clear dd d e' ee'. move => d1 v1 R1 dv. simpl.
|
||||
rewrite -> (fmon_eq_elim UR_id d1). simpl. move => A.
|
||||
destruct (kleisliValVal A) as [d2 [B Q]]. have e1 := vinj Q. clear Q A.
|
||||
destruct (kleisliValVal B) as [d3 [C Q]]. clear B. destruct (kleisliValVal C) as [d4 [C0 Q0]].
|
||||
clear C. destruct (kleisliValVal C0) as [d5 [C1 Q1]]. clear C0. have X':= (X (d1,v1) R1 _ C1).
|
||||
specialize (P _ v1 X'). have Q2:=vinj Q1. clear Q1. rewrite <- Q2 in Q0. clear Q2 d4.
|
||||
destruct (kleisliValVal Q) as [d4 [P' Q']]. clear Q. have e2:=vinj Q'. clear Q'.
|
||||
have ee:kleisli (eta << Unroll) (f' (Roll d5)) =-= Val (Unroll d3). rewrite -> Q0. by rewrite kleisliVal.
|
||||
specialize (P _ ee). case: P => v2 Ev. exists v2. case: Ev => Ev P. split. apply Ev.
|
||||
specialize (X _ P _ P'). rewrite <- e2 in e1. clear e2 d2. have e2:=fmon_eq_elim UR_id d4.
|
||||
rewrite -> e2 in e1. simpl in e1. clear e2. have ea:= (frespect S).
|
||||
have eb:S (dv, v2) =-= S (d4, v2). apply frespect. split ; last by []. by rewrite -> e1.
|
||||
by apply (proj2 eb).
|
||||
Qed.
|
||||
|
||||
Lemma imageD_eq: forall f g X Y, f =-= g -> imageD f X Y -> imageD g X Y.
|
||||
intros f g X Y fg. unfold imageD. intros C. intros x ; case x ; clear x.
|
||||
intros d v. simpl. specialize (C (d,v)). intros Xd. specialize (C Xd).
|
||||
simpl in C. intros dd gd. assert (ff := fmon_eq_elim fg d). rewrite <- ff in gd.
|
||||
specialize (C _ gd). apply C.
|
||||
Qed.
|
||||
|
||||
Lemma imageD_eqS : forall (X Y Z W:RelKind =-> Props) f, @Ole (set_ordType _) Z X ->
|
||||
@Ole (set_ordType RelKind) Y W -> imageD f X Y -> imageD f Z W.
|
||||
move => X Y Z W f l x y. case => e v Ze.
|
||||
specialize (l (e,v) Ze). specialize (y (e,v)). simpl in y. move => fe E.
|
||||
specialize (y l _ E). by apply x.
|
||||
Qed.
|
||||
|
||||
Lemma Delta_l: Delta =-= Deltam.
|
||||
split. apply (proj1 (RelV_least (proj1 RelV_fixedD) (proj1 (RelV_fixedDm)))).
|
||||
case_eq Delta. intros D AD Deq. case_eq Deltam. intros Dm ADm Dmeq. simpl.
|
||||
intros dv. case dv. clear dv. intros d v. unfold RAdm in AD.
|
||||
intros Dmdv. assert (imageD (eta << Unroll << Id << Roll) Dm D).
|
||||
assert (AA:=AD Dm).
|
||||
assert (admissible (fun (X:DInf -=> DInf _BOT) => imageD (((kleisli (eta << Unroll) << X) << Roll)) Dm D /\
|
||||
X <= eta)) as Ad.
|
||||
unfold admissible. intros c C1.
|
||||
split. assert (C := fun n => proj1 (C1 n)). clear C1.
|
||||
have a:= comp_eq_compat (comp_lub_eq (kleisli (eta << Unroll)) c) (tset_refl Roll).
|
||||
rewrite -> lub_comp in a.
|
||||
apply: (imageD_eq (Oeq_sym a)). clear a.
|
||||
apply AD. simpl. move => n. by apply C.
|
||||
apply lub_le => n. by apply (proj2 (C1 n)).
|
||||
have ee:(((eta << Unroll) << Id) << Roll) =-= kleisli (eta << Unroll) << (FIXP delta) << Roll.
|
||||
rewrite <- id_min. rewrite kleisli_eta_com. by rewrite comp_idR.
|
||||
refine (imageD_eq (Oeq_sym ee) _). clear ee.
|
||||
rewrite FIXP_simpl.
|
||||
|
||||
refine (proj1 (fixp_ind Ad _ _)). unfold imageD. split.
|
||||
intros x ; case x ; clear x. simpl.
|
||||
intros dd vv. intros ddC d0. repeat (rewrite fcont_comp_simpl).
|
||||
move => F. rewrite -> kleisli_bot in F. by case: (@PBot_incon _ _ (proj2 F)).
|
||||
by apply: leastP.
|
||||
|
||||
intros f imaf. split ; last by rewrite -> (fmonotonic ( delta) (proj2 imaf)) ; rewrite -> delta_eta.
|
||||
|
||||
destruct imaf as [imaf fless].
|
||||
have aa:=comp_eq_compat (comp_eq_compat (tset_refl (kleisli (eta << Unroll))) (delta_simpl f)) (tset_refl Roll).
|
||||
repeat rewrite -> comp_assoc in aa. rewrite -> kleisli_eta_com in aa.
|
||||
do 3 rewrite <- (comp_assoc Roll) in aa. rewrite -> UR_id in aa. simpl. rewrite -> comp_idR in aa.
|
||||
apply (imageD_eq (Oeq_sym aa)).
|
||||
|
||||
assert (X1:=RelV_fixedD).
|
||||
assert (X2:=RelV_fixedDm).
|
||||
rewrite Deq in X1, X2. rewrite Dmeq in X1, X2.
|
||||
assert (D =-= RelV Dm D) by auto.
|
||||
assert (Dm =-= RelV D Dm) by auto.
|
||||
clear X1 X2.
|
||||
|
||||
apply (imageD_eqS (proj1 H0) (proj2 H)).
|
||||
have e0:((kleisli (eta << Roll) <<
|
||||
((kleisli (eta << Unroll) << f) << Roll)) << Unroll) =-= f.
|
||||
repeat rewrite comp_assoc. rewrite <- comp_assoc. rewrite RU_id. rewrite comp_idR.
|
||||
rewrite <- kleisli_comp2. rewrite <- comp_assoc. rewrite RU_id. rewrite comp_idR.
|
||||
rewrite kleisli_unit. by rewrite comp_idL.
|
||||
apply: (imageD_eq _ (RelV_F_action imaf)). rewrite -> e0. repeat rewrite <- comp_assoc.
|
||||
rewrite UR_id. by rewrite comp_idR.
|
||||
|
||||
unfold imageD in H. specialize (H _ Dmdv). simpl in H.
|
||||
specialize (H d). apply H. unfold id. apply: (fmon_stable eta).
|
||||
by rewrite -> (fmon_eq_elim UR_id d).
|
||||
Qed.
|
||||
|
||||
Definition LR : RelKind =-> Props := match Delta with exist D _ => D end.
|
||||
|
||||
Lemma LR_fix: LR =-= (RelV LR LR) /\ RAdm LR.
|
||||
unfold LR.
|
||||
assert (Dfix:=RelV_fixedD). assert (monotonic RelVm) by auto.
|
||||
assert (M:= H (Delta,Delta) (Deltam,Delta)).
|
||||
specialize (H (Deltam,Delta) (Delta,Delta)).
|
||||
assert (xx:=Delta_l). split.
|
||||
assert (((Deltam, Delta) : (prod_clatType RelAdmop RelAdm)) <= (Delta, Delta)). destruct xx as [X0 X1].
|
||||
split. auto. auto. specialize (H H0).
|
||||
assert (((Delta, Delta) : (prod_clatType RelAdmop RelAdm)) <= (Deltam, Delta)). destruct xx as [X0 X1].
|
||||
split ; auto.
|
||||
specialize (M H1). clear H0 H1.
|
||||
generalize H M Dfix xx. clear Dfix xx H M.
|
||||
case Delta. intros D Rd.
|
||||
case Deltam. intros Dm RDm.
|
||||
intros H M Dfix yy.
|
||||
assert (D =-= RelV Dm D) by auto.
|
||||
assert (Dm =-= D) by auto. clear yy Dfix.
|
||||
case: H. case: M. simpl. unfold Ole. simpl. rewrite {1 3} / Ole. simpl.
|
||||
move => R0 R1 R2 R3. split => A.
|
||||
- apply R3. by apply (proj1 H0).
|
||||
- apply (proj2 H0). by apply R1.
|
||||
case Delta. simpl. by auto.
|
||||
Qed.
|
||||
|
||||
Definition ELR := fun ld e, forall d, ld =-= Val d -> exists v, exists ev:e =>> v, (RelV LR LR) (d,v).
|
||||
|
||||
Lemma RelV_simpl: forall R S d v, (RelV R S) (d,v) =
|
||||
match v with
|
||||
| INT m => d =-= inl(* _ *) m
|
||||
| VAR _ => False
|
||||
| LAMBDA e => exists f, d =-= inr (* _*) f /\
|
||||
forall (d1:VInf) (v1:Value O), R (d1,v1) ->
|
||||
forall dv, kleisli (eta << Unroll) (f (Roll d1)) =-= Val dv ->
|
||||
exists v2, subExp (cons v1 (@idSub _)) e =>> v2 /\ S (dv,v2)
|
||||
end. intros R S d v. simpl. case v ; auto.
|
||||
Qed.
|
||||
|
||||
Fixpoint LRsubst E : SemEnv E -> unii.Sub E 0 -> Prop :=
|
||||
match E with
|
||||
| O => fun d m => True
|
||||
| S n' => fun d m => LR (snd d, hd m) /\ LRsubst (fst d) (tl m)
|
||||
end.
|
||||
|
||||
Add Parametric Morphism (O0 O1:cpoType) : (@pair O0 O1)
|
||||
with signature (@tset_eq _ : O0 -> O0 -> Prop) ==> (@tset_eq _ : O1 -> O1 -> Prop) ==> (@tset_eq (O0 * O1))
|
||||
as pair_eq_compat.
|
||||
move => x y e x' y' e'. split. by rewrite -> e ; rewrite -> e'.
|
||||
by rewrite <- e ; rewrite <- e'.
|
||||
Qed.
|
||||
|
||||
(*=Fundamental *)
|
||||
Theorem FundamentalTheorem n :
|
||||
(forall (v:Value n) sl d, LRsubst d sl -> (RelV LR LR) (SemVal v d, subVal sl v)) /\
|
||||
forall (e:Exp n) sl d, LRsubst d sl -> ELR (SemExp e d) (subExp sl e).
|
||||
(*=End *)
|
||||
rewrite (lock LR).
|
||||
move: n ; apply ExpValue_ind.
|
||||
(* VAR *)
|
||||
- intros n v sl d C.
|
||||
dependent induction v.
|
||||
+ destruct C as [Ch Ct]. fold LRsubst in Ct. rewrite -lock.
|
||||
have X:= ((proj1 LR_fix) _). by apply (proj1 (X _)).
|
||||
+ destruct C as [Ch Ct]. fold LRsubst in Ct. rewrite -lock. rewrite (consEta sl). rewrite -lock in IHv. by apply (IHv _ _ Ct).
|
||||
(* INT *)
|
||||
- by [].
|
||||
(* LAMBDA *)
|
||||
- move => n e IH s d C. simpl SemVal. SimplMap.
|
||||
exists ((exp_fun ((kleisli (eta << Roll) << SemExp e) << Id >< Unroll)) d).
|
||||
split ; first by [].
|
||||
+ intros d1 v1 c1 dv. simpl.
|
||||
move => X.
|
||||
have Y:=Oeq_trans (fmon_eq_elim (kleisli_comp2 _ _) _) X. clear X.
|
||||
case: (kleisliValVal Y) => v [se P]. rewrite <- comp_assoc in P. rewrite -> UR_id in P.
|
||||
rewrite -> P in se. clear P v Y.
|
||||
specialize (IH (cons v1 s) (d,d1)).
|
||||
have: (LRsubst ((d,d1):SemEnv (n.+1)) (cons v1 s)). split.
|
||||
* rewrite -lock in c1. by apply c1.
|
||||
* simpl. rewrite tlCons. by apply C.
|
||||
move => H. specialize (IH H).
|
||||
unfold ELR in IH. specialize (IH dv).
|
||||
have S: (SemExp e (d, d1) =-= Val dv). rewrite <- se.
|
||||
by rewrite -> (pair_eq_compat (tset_refl _) (fmon_eq_elim UR_id d1)).
|
||||
rewrite -> (pair_eq_compat (tset_refl _) (fmon_eq_elim UR_id d1)) in se. specialize (IH se).
|
||||
case: IH => v [ev evH]. exists v. split. rewrite <- (proj2 (applyComposeSub _)). rewrite composeSingleSub. by apply ev. rewrite -lock. by apply (proj2 (proj1 LR_fix _)).
|
||||
|
||||
(* VAL *)
|
||||
- move => n v IH s d l d' H.
|
||||
specialize (IH s d l). exists (subVal s v). exists (e_Val (subVal s v)). rewrite -lock in IH.
|
||||
simpl in H. have e:=vinj H. clear H.
|
||||
refine (proj1 (frespect (RelV LR LR) _) IH). by apply NSetoid.pair_eq_compat.
|
||||
(* APP *)
|
||||
- move => n v0 IH0 v1 IH1 s d l d' e. SimplMap.
|
||||
simpl in e. specialize (IH0 s d l). specialize (IH1 s d l).
|
||||
simpl in IH0. case: (subVal s v0) IH0 ; first by [].
|
||||
+ move=> m ee. rewrite -> ee in e. rewrite -> SUM_fun_simplx in e. simpl in e. rewrite -> kleisli_bot in e.
|
||||
by case: (PBot_incon (proj2 e)).
|
||||
+ move => e'. case => f [e0 P]. rewrite -> e0 in e. rewrite -> SUM_fun_simplx in e.
|
||||
rewrite -lock in IH1. have IH:= proj2 (proj1 LR_fix _) IH1. clear IH1.
|
||||
rewrite -lock in P.
|
||||
specialize (P _ _ IH _ e). case: P => v. case => ev lr.
|
||||
exists v. exists (e_App ev). by apply (proj1 (proj1 LR_fix _)).
|
||||
(* LET *)
|
||||
- move => n e0 IH0 e1 IH1 s d l. simpl SemExp. SimplMap.
|
||||
specialize (IH0 s d l). simpl.
|
||||
move => d' X.
|
||||
case: (KLEISLIR_ValVal X). clear X.
|
||||
move => d0 [D0 D1]. specialize (IH0 d0 D0). case: IH0 => v0 [ev0 r0].
|
||||
specialize (IH1 (cons v0 s) (d,d0)). have rr:=proj2 (proj1 LR_fix _) r0. clear r0.
|
||||
unfold LRsubst in IH1. fold LRsubst in IH1. rewrite hdCons tlCons in IH1.
|
||||
specialize (IH1 (conj rr l)). unfold ELR in IH1. specialize (IH1 d' D1). destruct IH1 as [v [ev H]].
|
||||
exists v.
|
||||
split ; last by apply H.
|
||||
refine (e_Let ev0 _). rewrite <- (proj2 (applyComposeSub _)). rewrite composeSingleSub. by apply ev.
|
||||
(* IFZ *)
|
||||
- move => n v0 IH0 e1 IH1 e2 IH2 s d l d' X.
|
||||
specialize (IH0 s d l). simpl in IH0. SimplMap.
|
||||
simpl in X.
|
||||
case: (subVal s v0) IH0 ; first by [].
|
||||
+ move => ii ee. simpl SemExp in X. simpl in X. rewrite -> ee in X. rewrite -> SUM_fun_simplx in X. unfold id in X.
|
||||
specialize (IH1 s d l d'). specialize (IH2 s d l d').
|
||||
case: ii X ee.
|
||||
* move => Y ee. simpl in Y. rewrite -> SUM_fun_simplx in Y. simpl in Y.
|
||||
specialize (IH1 Y). case: IH1 => v [ev rv]. exists v. split. apply (e_Ifz1 _ ev). by apply rv.
|
||||
* move => ii Y ee. simpl in Y. rewrite -> SUM_fun_simplx in Y. simpl in Y.
|
||||
specialize (IH2 Y). case: IH2 => v [ev rv]. exists v. split. simpl. apply (e_Ifz2 _ _). apply ev. by apply rv.
|
||||
+ move => e. case => f [F _]. rewrite -> F in X. rewrite -> SUM_fun_simplx in X.
|
||||
simpl in X. by case: (PBot_incon (proj2 X)).
|
||||
(* OP *)
|
||||
- move => n op v0 IH0 v1 IH1 s d l d' X. simpl in X.
|
||||
specialize (IH0 s d l). specialize (IH1 s d l). simpl in IH0. simpl in IH1. SimplMap.
|
||||
case: (kleisliValVal X) => dd [Y Z]. clear X.
|
||||
case: (subVal s v0) IH0 IH1 ; first by [].
|
||||
+ move => i0 ee. case: (subVal s v1) ; first by [].
|
||||
* move => i1 ee'. rewrite -> ee in Y. rewrite -> ee' in Y. clear ee ee'.
|
||||
do 2 rewrite -> SUM_fun_simplx in Y. simpl in Y. unfold Smash in Y.
|
||||
rewrite -> Operator2_simpl in Y. have e:=vinj Y. clear Y.
|
||||
rewrite <- e in Z. clear e dd. exists (@INT 0 (op i0 i1)). exists (e_Op op i0 i1).
|
||||
simpl in Z. simpl. by rewrite <- (vinj Z).
|
||||
* move => e1. case => f [F _]. rewrite -> F in Y. rewrite -> ee in Y.
|
||||
do 2 rewrite -> SUM_fun_simplx in Y. simpl in Y. unfold Smash in Y.
|
||||
rewrite -> Operator2_strictR in Y. by case: (PBot_incon (proj2 Y)).
|
||||
+ move => e1. case => f [F _]. rewrite -> F in Y. do 2 rewrite -> SUM_fun_simplx in Y. simpl in Y.
|
||||
unfold Smash in Y. rewrite -> Operator2_strictL in Y.
|
||||
by case: (PBot_incon (proj2 Y)).
|
||||
Qed.
|
||||
|
||||
(*=Adequate *)
|
||||
Corollary Adequacy (e:Exp O) d : SemExp e tt =-= Val d -> exists v, e =>> v.
|
||||
(*=End *)
|
||||
move => S.
|
||||
assert (xx:=proj2 (FundamentalTheorem _) e (@idSub _) tt).
|
||||
simpl in xx. unfold ELR in xx. specialize (xx I _ S). destruct xx as [v P]. exists v.
|
||||
have ee:subExp (@idSub _) e = e by apply (proj2 (Map.applyId _ _)). rewrite ee in P. by apply P.
|
||||
Qed.
|
||||
|
29
papers/domains-2012/uniiop.v
Normal file
29
papers/domains-2012/uniiop.v
Normal file
|
@ -0,0 +1,29 @@
|
|||
(**********************************************************************************
|
||||
* uniiop.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
Require Export unii.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
|
||||
Reserved Notation "e '=>>' v" (at level 75).
|
||||
|
||||
(*=Evaluation *)
|
||||
Inductive Evaluation : Exp O -> Value O -> Prop :=
|
||||
| e_Val : forall v, VAL v =>> v
|
||||
| e_App : forall e1 v2 v, subExp [v2] e1 =>> v -> APP (LAMBDA e1) v2 =>> v
|
||||
| e_Let : forall e1 v1 e2 v2, e1 =>> v1 -> subExp [v1] e2 =>> v2 -> LET e1 e2 =>> v2
|
||||
| e_Ifz1 : forall e1 e2 v1, e1 =>> v1 -> IFZ (INT 0) e1 e2 =>> v1
|
||||
| e_Ifz2 : forall e1 e2 v2 n, e2 =>> v2 -> IFZ (INT (S n)) e1 e2 =>> v2
|
||||
| e_Op : forall op n1 n2, OP op (INT n1) (INT n2) =>> INT (op n1 n2)
|
||||
where "e =>> v" := (Evaluation e v).
|
||||
|
||||
(*=End *)
|
||||
|
421
papers/domains-2012/uniirec.v
Normal file
421
papers/domains-2012/uniirec.v
Normal file
|
@ -0,0 +1,421 @@
|
|||
(**********************************************************************************
|
||||
* uniirec.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* Construction of recursive domain for interpreting unityped lambda calculus *)
|
||||
|
||||
Require Export PredomAll.
|
||||
Require Import PredomRec.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
(*=kcpoCat *)
|
||||
Lemma kcpoCatAxiom : @Category.axiom cpoType
|
||||
(fun X Y => exp_cppoType X (liftCppoType Y)) (fun X Y Z f g => kleisli f << g) (@eta).
|
||||
(*CLEAR*)
|
||||
split ; last split ; last split.
|
||||
- move => D0 D1 f. simpl. rewrite kleisli_unit. by rewrite comp_idL.
|
||||
- move => D0 D1 f. simpl. by rewrite kleisli_eta_com.
|
||||
- move => D0 D1 D2 D3 f g h. simpl. rewrite <- kleisli_comp. by rewrite comp_assoc.
|
||||
- move => D0 D1 D2 f f' g g'. simpl. move => e e'. rewrite e. by rewrite e'.
|
||||
Qed.
|
||||
(*CLEARED*)Canonical Structure kcpoCatMixin := CatMixin kcpoCatAxiom.
|
||||
Canonical Structure kcpoCatType := Eval hnf in CatType kcpoCatMixin.
|
||||
(*=End *)
|
||||
|
||||
Module Type RecDom.
|
||||
Variable DInf : cpoType.
|
||||
Definition VInf := discrete_cpoType nat + (DInf -=> DInf _BOT).
|
||||
Variable Roll : VInf =-> DInf.
|
||||
Variable Unroll : DInf =-> VInf.
|
||||
|
||||
Variable RU_id : Roll << Unroll =-= Id.
|
||||
Variable UR_id : Unroll << Roll =-= Id.
|
||||
|
||||
Variable delta : (DInf -=> DInf _BOT) =-> (DInf -=> DInf _BOT).
|
||||
Variable delta_simpl : forall e, delta e =-= eta << Roll << ([| in1,
|
||||
(in2 <<
|
||||
((exp_fun (CCOMP DInf (DInf _BOT) (DInf _BOT):cpoCatType _ _) (kleisli e) : cpoCatType _ _) <<
|
||||
((exp_fun
|
||||
((CCOMP DInf (DInf _BOT) (DInf _BOT)) << SWAP) e :cpoCatType _ _) << KLEISLI))) |]) << Unroll.
|
||||
|
||||
Variable delta_eta : delta eta =-= eta.
|
||||
Variable id_min : eta =-= @FIXP _ delta.
|
||||
|
||||
End RecDom.
|
||||
|
||||
Module RD : RecDom.
|
||||
|
||||
Lemma kcpoTerminalAxiom : CatTerminal.axiom (Zero: kcpoCatType).
|
||||
simpl. move => D x y. split.
|
||||
move => i. simpl. apply: DLless_cond. by case.
|
||||
move => i. simpl. apply:DLless_cond. by case.
|
||||
Qed.
|
||||
|
||||
Canonical Structure kcpoTermincalCatMixin :=
|
||||
@terminalCatMixin kcpoCatType (Zero: kcpoCatType)
|
||||
(fun X => const _ (PBot: (liftCpoPointedType Zero))) kcpoTerminalAxiom.
|
||||
Canonical Structure kcpoTerminalCat := Eval hnf in @terminalCatType kcpoCatType kcpoTermincalCatMixin.
|
||||
|
||||
Lemma kcpo_comp_eq (X Y Z : cpoType) m m' : ((CCOMP X (Y _BOT) (Z _BOT)) << KLEISLI >< Id) (m,m') =-= Category.tcomp kcpoCatMixin m m'.
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition kcpoBaseCatMixin := CppoECatMixin kcpoTermincalCatMixin kcpo_comp_eq.
|
||||
|
||||
(*=kcpoBaseCat *)
|
||||
Canonical Structure kcpoBaseCatType := Eval hnf in CppoECatType kcpoBaseCatMixin.
|
||||
Lemma leftss : (forall (X Y Z : kcpoBaseCatType) (f : kcpoBaseCatType X Y),
|
||||
(PBot:kcpoCatType _ _) << f =-= (PBot: X =-> Z)).
|
||||
(*=End *)
|
||||
move => X Y Z f. apply: fmon_eq_intro.
|
||||
move => x. split ; last by apply: leastP.
|
||||
apply: DLless_cond.
|
||||
move => z. move => A. case: (kleisliValVal A) => y [_ F]. by case: (PBot_incon_eq (Oeq_sym F)).
|
||||
Qed.
|
||||
|
||||
Definition ProjSet (T:Tower kcpoBaseCatType) := fun (d:prodi_cpoType (fun n => tobjects T n _BOT)) => forall n,
|
||||
PROJ (fun n => tobjects T n _BOT) n d =-=
|
||||
kleisli (tmorphisms T n) (PROJ (fun n => tobjects T n _BOT) (S n) d) /\
|
||||
exists n, exists e, PROJ (fun n => tobjects T n _BOT) n d =-= Val e.
|
||||
|
||||
Lemma ProjSet_inclusive T : admissible (@ProjSet T).
|
||||
unfold ProjSet. unfold admissible.
|
||||
intros c C n.
|
||||
split. do 3 rewrite -> lub_comp_eq.
|
||||
refine (lub_eq_compat _).
|
||||
refine (fmon_eq_intro _).
|
||||
intros m. simpl. specialize (C m n). by apply (proj1 C).
|
||||
specialize (C 0 0). destruct C as [_ C]. clear n.
|
||||
destruct C as [n [e P]].
|
||||
exists n.
|
||||
assert (forall n, continuous ((PROJ (fun n0 : nat => tobjects T n0 _BOT) n))) as Cp by auto.
|
||||
assert (PROJ (fun n : nat => tobjects T n _BOT) n (c 0) <= PROJ (fun n : nat => tobjects T n _BOT) n (lub c)) as L by
|
||||
(apply: fmonotonic ; auto).
|
||||
rewrite -> P in L.
|
||||
destruct (DLle_Val_exists_eq L) as [dn [Y X]].
|
||||
exists dn. by apply Y.
|
||||
Qed.
|
||||
|
||||
Definition kcpoCone (T:Tower kcpoBaseCatType) : Cone T.
|
||||
exists (sub_cpoType (@ProjSet_inclusive T)) (fun i:nat => PROJ _ i << Forget (@ProjSet_inclusive T)).
|
||||
move => i. apply: fmon_eq_intro. case => d Pd.
|
||||
by apply (Oeq_sym (proj1 (Pd i))).
|
||||
Defined.
|
||||
|
||||
Implicit Arguments InheritFun [D E P].
|
||||
|
||||
Lemma retract_total D E (f:D =-> E _BOT) (g:E =-> D _BOT) : kleisli f << g =-= eta -> total g.
|
||||
unfold total. move => X d. have X':=fmon_eq_elim X d.
|
||||
case: (kleisliValVal X'). move => e [Y _]. exists e. by apply Y.
|
||||
Qed.
|
||||
|
||||
Lemma xx (T:Tower kcpoBaseCatType) i : (forall d : tobjects T i, ProjSet (PRODI_fun (t_nm T i) d)).
|
||||
move => d n. split. simpl.
|
||||
by rewrite -> (fmon_eq_elim (t_nmProjection T i n) d).
|
||||
exists i. exists d. simpl. by rewrite t_nn_ID.
|
||||
Qed.
|
||||
|
||||
Definition kcpoCocone (T:Tower kcpoBaseCatType) : CoCone T.
|
||||
exists (sub_cpoType (@ProjSet_inclusive T)) (fun i => eta << @InheritFun _ _ _ (@ProjSet_inclusive T) (PRODI_fun (t_nm T i)) (@xx T i)).
|
||||
move => i. rewrite {1} /Category.comp. simpl. apply: fmon_eq_intro => d. split.
|
||||
- apply: DLless_cond. case => x Px C. case: (kleisliValVal C). clear C.
|
||||
move => y [md X].
|
||||
apply Ole_trans with (kleisli (eta << InheritFun (@ProjSet_inclusive T) (PRODI_fun (t_nm T i.+1)) (@xx T _)) (Val y)) ;
|
||||
first by rewrite <- md.
|
||||
rewrite kleisliVal. rewrite -> X. apply: (fmonotonic (@eta_m _)). unfold Ole. simpl.
|
||||
move => n. simpl. have Y:=vinj X.
|
||||
case: (fmon_stable (Forget (@ProjSet_inclusive T)) Y). clear Y. simpl. move => Y Y'.
|
||||
specialize (Y' n). rewrite -> Y'. rewrite -> (fmon_eq_elim (t_nmEmbedding T i n) d). simpl.
|
||||
rewrite -> md. by rewrite kleisliVal.
|
||||
- case: (retract_total (proj1 (teppair T i)) d). move => x e.
|
||||
apply Ole_trans with (y:=kleisli (eta << InheritFun (@ProjSet_inclusive T) (PRODI_fun (t_nm T i.+1)) (@xx T _)) (Val x)) ;
|
||||
last by rewrite <- e.
|
||||
rewrite kleisliVal. apply: DLle_leVal. move => n. simpl.
|
||||
apply Ole_trans with (y:=(kleisli (t_nm T i.+1 n) (Val x))) ; last by rewrite kleisliVal.
|
||||
rewrite <- e. by apply (proj1 (fmon_eq_elim (t_nmEmbedding T i n) d)).
|
||||
Defined.
|
||||
|
||||
Lemma limit_def (T:Tower kcpoBaseCatType) (C:Cone T) d n e' : mcone C n d =-= Val e' ->
|
||||
exists e, lub (chainPE (kcpoCocone T) C) d =-= Val e.
|
||||
move => X. simpl.
|
||||
have aa:exists e, (fcont_app (chainPE (kcpoCocone T) C) d) n =-= Val e.
|
||||
exists (@InheritFun _ _ _ (@ProjSet_inclusive T) (PRODI_fun (t_nm T n)) (@xx T n) e').
|
||||
apply (@Oeq_trans _ _ (kleisli (eta << InheritFun (@ProjSet_inclusive T) (PRODI_fun (t_nm T n)) (@xx T _)) (mcone C n d))) ; first by [].
|
||||
rewrite -> X. by rewrite kleisliVal.
|
||||
case: aa => e aa. case: (chainVallubnVal 1 aa) => x bb. exists x. by apply bb.
|
||||
Qed.
|
||||
|
||||
(*=kcpoLimit *)
|
||||
Definition kcpoLimit (T:Tower kcpoBaseCatType) : Limit T.
|
||||
(*=End *)
|
||||
exists (kcpoCone T) (fun C : Cone T => lub (chainPE (@kcpoCocone T) C)).
|
||||
move => C n. simpl. split.
|
||||
- apply: (Ole_trans _ (comp_le_compat (Ole_refl _) (le_lub (chainPE (kcpoCocone T) C) n))).
|
||||
simpl. rewrite {1} /Category.comp. simpl. rewrite comp_assoc. rewrite <- kleisli_comp2.
|
||||
rewrite <- comp_assoc. rewrite -> ForgetInherit. rewrite prodi_fun_pi. rewrite t_nn_ID. rewrite kleisli_unit.
|
||||
by rewrite comp_idL. simpl.
|
||||
rewrite {1} / Category.comp. simpl.
|
||||
refine (Ole_trans (Oeq_le (PredomCore.comp_lub_eq _ (chainPE (kcpoCocone T) C))) _).
|
||||
rewrite (lub_lift_left _ n). apply: lub_le => i. simpl. rewrite comp_assoc.
|
||||
rewrite <- (kleisli_comp2 (InheritFun (@ProjSet_inclusive T) (PRODI_fun (t_nm T (n + i))) (@xx T _))
|
||||
(PROJ (fun n0 : nat => tobjects T n0 _BOT) n << Forget (@ProjSet_inclusive T))).
|
||||
rewrite <- comp_assoc. rewrite ForgetInherit. rewrite prodi_fun_pi. by apply (proj2 ((coneCom_l C (leq_addr i n)))).
|
||||
- move => C h X. apply: fmon_eq_intro => d. simpl in h. split.
|
||||
+ apply: DLless_cond. case => x Px E. case: (proj2 (Px 0)) => n. case => y Py. rewrite -> E.
|
||||
have A:=(fmon_eq_elim (X n) d). have AA:=tset_trans A (fmon_stable (kleisli _) E). clear A.
|
||||
have A:=tset_trans AA (kleisliVal _ _). clear AA. simpl in A. rewrite -> Py in A.
|
||||
case: (limit_def A) => lc e. rewrite -> e. apply: DLle_leVal. case: lc e => lc Plc e. unfold Ole. simpl.
|
||||
move => i. specialize (X i). have Xi:=fmon_eq_elim X d.
|
||||
have Xii: (mcone C i) d =-= (kleisli (PROJ _ i << Forget (@ProjSet_inclusive T)) ( h d)) by apply Xi.
|
||||
rewrite -> E in Xii. rewrite -> kleisliVal in Xii. simpl in Xii.
|
||||
rewrite <- Xii. clear Xi Xii.
|
||||
simpl in e. have aa := Ole_trans (le_lub _ i) (proj1 e). clear e h X E. simpl in aa.
|
||||
have bb:kleisli (eta << (@InheritFun _ _ _ (@ProjSet_inclusive T) (PRODI_fun (t_nm T (i))) (@xx T (i))))
|
||||
((mcone C i) d) <= Val (exist (fun x : forall i : nat, Stream (tobjects T i) => ProjSet x)
|
||||
lc Plc) by apply aa. clear aa.
|
||||
apply: DLless_cond => di X. rewrite -> X in bb. rewrite -> kleisliVal in bb. rewrite -> X.
|
||||
have aa:=vleinj bb. clear bb. unfold Ole in aa. simpl in aa. specialize (aa i). simpl in aa.
|
||||
rewrite <- aa. by rewrite -> (fmon_eq_elim (t_nn_ID T i) di).
|
||||
+ simpl. apply: lub_le => n. specialize (X n). have Y:=fmon_eq_elim X d. clear X.
|
||||
simpl mcone in Y. simpl. apply Ole_trans with (y:=kleisli (eta << (@InheritFun _ _ _ (@ProjSet_inclusive T) (PRODI_fun (t_nm T n)) (@xx T n))) ( (mcone C n) d)) ; first by [].
|
||||
rewrite -> Y.
|
||||
apply Ole_trans with (y:=(kleisli (eta << InheritFun (@ProjSet_inclusive T) (PRODI_fun (t_nm T n)) (@xx T _)) <<
|
||||
kleisli (PROJ (fun n0 : nat => tobjects T n0 _BOT) n << Forget (@ProjSet_inclusive T)))
|
||||
( h d)) ; first by [].
|
||||
apply: DLless_cond. move => aa X. rewrite -> X. case: (kleisliValVal X) => b [P Q]. clear X.
|
||||
case: (kleisliValVal P) => hd [P' Q']. rewrite -> P'. apply: DLle_leVal.
|
||||
rewrite <- (vinj Q). clear P Q aa h d Y P'. unfold Ole. case: hd Q' => x Px Q.
|
||||
simpl. simpl in Q. move => i. simpl.
|
||||
case: (ltngtP n i).
|
||||
* move => l. have a:= comp_eq_compat (tset_refl (t_nm T n i)) (coneCom_l (kcpoCone T) (ltnW l)).
|
||||
rewrite -> comp_assoc in a. have yy:t_nm T n i << mcone (kcpoCone T) n <= mcone (kcpoCone T) i.
|
||||
rewrite -> a. rewrite -> (comp_le_compat (proj2 (t_nm_EP T (ltnW l))) (Ole_refl _)).
|
||||
by rewrite comp_idL.
|
||||
specialize (yy (exist _ x Px)). simpl in yy. rewrite -> Q in yy. rewrite -> kleisliVal in yy.
|
||||
by apply yy.
|
||||
* move => l. have a:= (proj2 (fmon_eq_elim (coneCom_l (kcpoCone T) (ltnW l)) (exist _ x Px))).
|
||||
simpl in a. have aa:(kleisli ( (t_nm T n i)) (x n)) <= (x i) by apply a.
|
||||
rewrite -> Q in aa. rewrite -> kleisliVal in aa. by apply aa.
|
||||
* move => e. rewrite <- e. clear i e. rewrite -> (proj1 (fmon_eq_elim (t_nn_ID T n) b)). by rewrite -> Q.
|
||||
Defined.
|
||||
|
||||
Lemma summ_mon (F G : BiFunctor kcpoBaseCatType)
|
||||
X Y Z W : monotonic (fun p => [|kleisli (eta << in1) << (morph F X Y Z W p : (ob F X Z) =-> (ob F Y W)),
|
||||
kleisli (eta << in2) << (morph G X Y Z W p : (ob G X Z) =-> (ob G Y W))|]).
|
||||
move => p p' l. simpl.
|
||||
unfold sum_fun. simpl. unfold in1. simpl. unfold in2. simpl.
|
||||
move => x. simpl. do 2 rewrite -> SUM_fun_simpl. case: x.
|
||||
- move => s. simpl. by rewrite -> l.
|
||||
- move => s. simpl. by rewrite -> l.
|
||||
Qed.
|
||||
|
||||
Definition summ (F G : BiFunctor kcpoBaseCatType) X Y Z W := Eval hnf in mk_fmono (@summ_mon F G X Y Z W).
|
||||
|
||||
Lemma sumc (F G : BiFunctor kcpoBaseCatType) X Y Z W : continuous (@summ F G X Y Z W).
|
||||
move => c. simpl. unfold sum_fun. simpl. move => x. simpl. rewrite -> SUM_fun_simpl. simpl.
|
||||
case:x ; simpl => s.
|
||||
- do 2 rewrite lub_comp_eq. simpl. apply lub_le_compat => i. simpl. unfold sum_fun. simpl. by rewrite SUM_fun_simpl.
|
||||
- do 2 rewrite lub_comp_eq. simpl. apply lub_le_compat => i. simpl. unfold sum_fun. simpl. by rewrite SUM_fun_simpl.
|
||||
Qed.
|
||||
|
||||
Definition sum_func (F G : BiFunctor kcpoBaseCatType) X Y Z W := Eval hnf in mk_fcont (@sumc F G X Y Z W).
|
||||
|
||||
Lemma sum_func_simpl F G X Y Z W x : @sum_func F G X Y Z W x = [|kleisli (eta << in1) << (morph F X Y Z W x : (ob F X Z) =-> (ob F Y W)),
|
||||
kleisli (eta << in2) << (morph G X Y Z W x : (ob G X Z) =-> (ob G Y W))|].
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition biSum (F G : BiFunctor kcpoBaseCatType) : BiFunctor kcpoBaseCatType.
|
||||
exists (fun X Y => (ob F X Y) + (ob G X Y)) (fun X Y Z W => @sum_func F G X Y Z W).
|
||||
move => T0 T1 T2 T3 T4 T5 f g h k. simpl.
|
||||
apply: (@sum_unique cpoSumCatType).
|
||||
- rewrite sum_fun_fst. rewrite {2} / Category.comp. simpl. rewrite <- comp_assoc.
|
||||
rewrite sum_fun_fst. rewrite comp_assoc. rewrite <- kleisli_comp2. rewrite sum_fun_fst.
|
||||
rewrite <- (comp_eq_compat (tset_refl (kleisli (eta << in1))) (@morph_comp _ F T0 T1 T2 T3 T4 T5 f g h k)).
|
||||
rewrite {6} /Category.comp. simpl. rewrite comp_assoc. by rewrite kleisli_comp.
|
||||
- rewrite sum_fun_snd. rewrite {2} / Category.comp. simpl. rewrite <- comp_assoc.
|
||||
rewrite sum_fun_snd. rewrite comp_assoc. rewrite <- kleisli_comp2. rewrite sum_fun_snd.
|
||||
rewrite <- (comp_eq_compat (tset_refl (kleisli (eta << in2))) (@morph_comp _ G T0 T1 T2 T3 T4 T5 f g h k)).
|
||||
rewrite {6} /Category.comp. simpl. rewrite comp_assoc. by rewrite kleisli_comp.
|
||||
- move => T0 T1. simpl. apply: (@sum_unique cpoSumCatType).
|
||||
+ simpl. rewrite sum_fun_fst. rewrite (comp_eq_compat (tset_refl (kleisli (eta << in1))) (morph_id F _ _)).
|
||||
by rewrite kleisli_eta_com.
|
||||
+ simpl. rewrite sum_fun_snd. rewrite (comp_eq_compat (tset_refl (kleisli (eta << in2))) (morph_id G _ _)).
|
||||
by rewrite kleisli_eta_com.
|
||||
Defined.
|
||||
|
||||
Lemma bifunm
|
||||
X Y Z W : monotonic (fun (p:@cppoMorph kcpoBaseCatType Y X * @cppoMorph kcpoBaseCatType Z W) =>
|
||||
eta << (exp_fun (CCOMP _ _ _ :cpoCatType _ _) (kleisli (snd p) : cpoCatType _ _)) << (exp_fun ((CCOMP _ _ _) << SWAP) (fst p)) << KLEISLI).
|
||||
move => p p' l f.
|
||||
simpl. apply: DLle_leVal. case: l => l l'. rewrite l. by rewrite -> (kleisli_le_compat l').
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism (D:cpoType) : (@Val D)
|
||||
with signature (@Ole D: D -> D -> Prop) ++> (@Ole (D _BOT))
|
||||
as Val_le_cpo_compat.
|
||||
intros.
|
||||
apply: DLle_leVal.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
|
||||
Lemma bifunc X Y Z W : continuous (mk_fmono (@bifunm X Y Z W)).
|
||||
move => c x. simpl.
|
||||
apply Ole_trans with (y:=eta (((KLEISLI (lub (pi2 << (c:natO =-> _))):cpoCatType _ _) <<
|
||||
exp_fun (CCOMP _ _ (_ _BOT):cpoCatType _ _)
|
||||
(kleisli x) (lub (pi1 << (c:natO =-> _)))))) ; first by [].
|
||||
do 2 rewrite lub_comp_eq. rewrite -> PredomCore.lub_comp_both.
|
||||
rewrite lub_comp_eq. by apply lub_le_compat => n.
|
||||
Qed.
|
||||
|
||||
Definition bi_fun (X Y Z W : kcpoBaseCatType) : (@cppoMorph kcpoBaseCatType Y X * cppoMorph Z W) =->
|
||||
(@cppoMorph kcpoBaseCatType (fcont_cpoType X (Z _BOT)) (fcont_cpoType Y (W _BOT)))
|
||||
:= Eval hnf in mk_fcont (@bifunc X Y Z W).
|
||||
|
||||
|
||||
Lemma bi_fun_simpl T0 T2 T4 T5 f g x : (bi_fun T0 T4 T2 T5) (f,g) x = Val (kleisli g << (kleisli x << f)).
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition biFun : BiFunctor kcpoBaseCatType.
|
||||
exists (fun X Y => fcont_cpoType X (Y _BOT)) (fun X Y Z W => @bi_fun X Y Z W).
|
||||
move => T0 T1 T2 T3 T4 T5 f g h k. apply: fmon_eq_intro => x.
|
||||
apply Oeq_trans with (y:=kleisli ((bi_fun T1 T4 T3 T5) (f, g)) ((bi_fun T0 T1 T2 T3) (h, k) x)) ; first by [].
|
||||
rewrite bi_fun_simpl. rewrite kleisliVal. rewrite bi_fun_simpl.
|
||||
apply Oeq_trans with (y:= (bi_fun T0 T4 T2 T5) (h << f, g << k) x) ; last by [].
|
||||
rewrite bi_fun_simpl. apply: (fmon_stable eta).
|
||||
rewrite <- kleisli_comp. rewrite <- kleisli_comp. rewrite {6 8} /Category.comp. simpl.
|
||||
rewrite <- kleisli_comp. by repeat rewrite comp_assoc.
|
||||
|
||||
move => X Y. apply: fmon_eq_intro => x. apply: (fmon_stable eta).
|
||||
simpl. rewrite kleisli_unit. rewrite comp_idL. by rewrite kleisli_eta_com.
|
||||
Defined.
|
||||
|
||||
Definition biVar : BiFunctor kcpoBaseCatType.
|
||||
exists (fun X Y => Y) (fun X Y Z W => pi2).
|
||||
by [].
|
||||
by [].
|
||||
Defined.
|
||||
|
||||
Definition biConst (D:kcpoBaseCatType) : BiFunctor kcpoBaseCatType.
|
||||
exists (fun (X Y:kcpoBaseCatType) => D) (fun (X Y Z W:kcpoBaseCatType) => const _ eta).
|
||||
move => T0 T1 T2 T3 T4 T5 f g h k. simpl. unfold Category.comp. simpl.
|
||||
rewrite kleisli_unit. by rewrite comp_idL.
|
||||
move => T0 T1. by [].
|
||||
Defined.
|
||||
|
||||
(*=FS *)
|
||||
Definition FS := biSum (biConst (discrete_cpoType nat)) biFun.
|
||||
(*=End *)
|
||||
(*=DInf *)
|
||||
Definition DInf : cpoType := @DInf kcpoBaseCatType kcpoLimit FS leftss.
|
||||
Definition VInf := (discrete_cpoType nat) + (DInf -=> DInf _BOT).
|
||||
Definition Fold : VInf =-> DInf _BOT := Fold kcpoLimit FS leftss.
|
||||
Definition Unfold : DInf =-> VInf _BOT := Unfold kcpoLimit FS leftss.
|
||||
(*=End *)
|
||||
Lemma FU_iso : kleisli Fold << Unfold =-= eta.
|
||||
by apply (FU_id kcpoLimit FS leftss).
|
||||
Qed.
|
||||
|
||||
Lemma UF_iso : kleisli Unfold << Fold =-= eta.
|
||||
by apply (UF_id kcpoLimit FS leftss).
|
||||
Qed.
|
||||
|
||||
Lemma ob X Y : ob FS X Y = discrete_cpoType nat + (X -=> (Y _BOT)).
|
||||
by simpl.
|
||||
Qed.
|
||||
|
||||
Lemma morph1 X Y Z W f g x : morph FS X Y Z W (f,g) (INL _ _ x) =-= Val (INL _ _ x).
|
||||
simpl. unfold sum_fun. simpl. unlock SUM_fun. simpl. by rewrite kleisliVal.
|
||||
Qed.
|
||||
|
||||
Lemma morph2 X Y Z W f g x : morph FS X Y Z W (f,g) (INR _ _ x) =-= Val (INR _ _ (kleisli g << (kleisli x << f))).
|
||||
simpl. unfold sum_fun. simpl. unlock SUM_fun. simpl; by rewrite kleisliVal.
|
||||
Qed.
|
||||
|
||||
(*=Delta *)
|
||||
Definition delta : (DInf -=> DInf _BOT) =-> (DInf -=> DInf _BOT) := delta kcpoLimit FS leftss.
|
||||
(*=End *)
|
||||
|
||||
Lemma eta_mono X Y (f g : X =-> Y) : eta << f =-= eta << g -> f =-= g.
|
||||
move => A.
|
||||
apply: fmon_eq_intro => x.
|
||||
have A':=fmon_eq_elim A x. by apply (vinj A').
|
||||
Qed.
|
||||
|
||||
(*=ROLL *)
|
||||
Lemma foldT : total Fold.
|
||||
(*CLEAR*)
|
||||
move => x. simpl.
|
||||
have X:=fmon_eq_elim UF_iso x. case: (kleisliValVal X). clear X. move => y [P Q]. exists y. by apply P.
|
||||
Qed.
|
||||
(*CLEARED*)Lemma unfoldT : total Unfold. (*CLEAR*)
|
||||
move => x. simpl.
|
||||
have X:=fmon_eq_elim FU_iso x. case: (kleisliValVal X). clear X. move => y [P Q]. exists y. by apply P.
|
||||
Qed. (*CLEARED*)
|
||||
Definition Roll : VInf =-> DInf := totalL foldT.
|
||||
Definition Unroll : DInf =-> VInf := totalL unfoldT.
|
||||
Lemma RU_id : Roll << Unroll =-= Id. (*CLEAR*)
|
||||
apply eta_mono.
|
||||
have X:=FU_iso.
|
||||
have A:eta << Roll =-= Fold by apply totalL_eta.
|
||||
rewrite <- A in X. clear A.
|
||||
have A:eta << Unroll =-= Unfold by apply totalL_eta.
|
||||
rewrite <- A in X. clear A.
|
||||
rewrite -> comp_assoc in X. rewrite -> kleisli_eta_com in X.
|
||||
rewrite <- comp_assoc in X. rewrite X. by rewrite comp_idR.
|
||||
Qed. (*CLEARED*)
|
||||
Lemma UR_id : Unroll << Roll =-= Id.
|
||||
(*=End *)
|
||||
(*=End *)
|
||||
apply eta_mono.
|
||||
have X:=UF_iso.
|
||||
have A:eta << Roll =-= Fold by apply totalL_eta.
|
||||
rewrite <- A in X. clear A.
|
||||
have A:eta << Unroll =-= Unfold by apply totalL_eta.
|
||||
rewrite <- A in X. clear A.
|
||||
rewrite -> comp_assoc in X. rewrite -> kleisli_eta_com in X.
|
||||
rewrite <- comp_assoc in X. rewrite X. by rewrite comp_idR.
|
||||
Qed.
|
||||
|
||||
Lemma delta_simpl (e:DInf =-> DInf _BOT) : delta e =-=
|
||||
eta << Roll << ([| in1,
|
||||
(in2 <<
|
||||
((exp_fun (CCOMP DInf (DInf _BOT) (DInf _BOT):cpoCatType _ _) (kleisli e) : cpoCatType _ _) <<
|
||||
((exp_fun
|
||||
((CCOMP DInf _ (DInf _BOT)) <<
|
||||
SWAP) e : cpoCatType _ _) << KLEISLI))) |]) << Unroll.
|
||||
rewrite (@delta_simpl _ kcpoLimit FS leftss e).
|
||||
fold Fold. fold Unfold. fold DInf. simpl. rewrite <- comp_assoc.
|
||||
rewrite {1 2} /Category.comp. simpl. have A:eta << Unroll =-= Unfold by apply totalL_eta.
|
||||
rewrite <- A. rewrite (comp_assoc Unroll eta). rewrite kleisli_eta_com.
|
||||
rewrite comp_assoc. apply: (comp_eq_compat _ (tset_refl Unroll)).
|
||||
have B:eta << Roll =-= Fold by apply totalL_eta.
|
||||
rewrite <- B. rewrite <- (comp_eq_compat (kleisli_eta_com (eta << Roll)) (tset_refl ([| _,_|]))).
|
||||
rewrite <- (comp_assoc _ eta). apply comp_eq_compat ; first by [].
|
||||
rewrite kleisli_eta_com.
|
||||
do 4 rewrite comp_assoc. rewrite kleisli_eta_com. simpl. apply: sum_unique.
|
||||
- rewrite sum_fun_fst. do 2 rewrite <- comp_assoc. by rewrite sum_fun_fst.
|
||||
- rewrite sum_fun_snd. repeat rewrite <- comp_assoc. by rewrite sum_fun_snd.
|
||||
Qed.
|
||||
|
||||
(*=minimal *)
|
||||
Lemma id_min : eta =-= FIXP delta.
|
||||
(*=End *)
|
||||
apply tset_sym. rewrite <- (id_min kcpoLimit FS leftss). fold delta.
|
||||
simpl. apply:fmon_eq_intro => n. simpl. apply lub_eq_compat. by apply fmon_eq_intro => m.
|
||||
Qed.
|
||||
|
||||
Lemma delta_eta : delta eta =-= eta.
|
||||
by apply (delta_id_id kcpoLimit FS leftss).
|
||||
Qed.
|
||||
|
||||
End RD.
|
121
papers/domains-2012/uniisem.v
Executable file
121
papers/domains-2012/uniisem.v
Executable file
|
@ -0,0 +1,121 @@
|
|||
(**********************************************************************************
|
||||
* uniisem.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* semantics of unityped lambda calculus in recursive domain *)
|
||||
|
||||
Require Import PredomAll.
|
||||
Require Import Fin unii uniirec.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
Export RD.
|
||||
|
||||
(*=SemEnv *)
|
||||
Fixpoint SemEnv E : cpoType := match E with O => One | S E => SemEnv E * VInf end.
|
||||
Fixpoint SemVar E (v : Fin E) : SemEnv E =-> VInf :=
|
||||
match v with
|
||||
| FinZ _ => pi2
|
||||
| FinS _ v => SemVar v << pi1
|
||||
end.
|
||||
(*=End *)
|
||||
|
||||
Canonical Structure nat_cpoType := Eval hnf in discrete_cpoType nat.
|
||||
Canonical Structure bool_cpoType := Eval hnf in discrete_cpoType bool.
|
||||
|
||||
Lemma zeroCase_mon : monotonic (fun (n:nat_cpoType) => match n with | O => @in1 _ (One:cpoType) _ tt | S m => @in2 _ _ (discrete_cpoType nat) m end).
|
||||
move => x y. case. move => e ; rewrite e. clear x e. by case: y.
|
||||
Qed.
|
||||
|
||||
Definition zeroCasem : ordCatType (discrete_cpoType nat) ((One:cpoType) + nat_cpoType) :=
|
||||
Eval hnf in mk_fmono zeroCase_mon.
|
||||
|
||||
Lemma zeroCase_cont : continuous zeroCasem.
|
||||
move => c. simpl. have e:lub c = c 0. by []. rewrite e. simpl.
|
||||
by apply: (Ole_trans _ (le_lub _ O)).
|
||||
Qed.
|
||||
|
||||
(*=zeroCase *)
|
||||
Definition zeroCase : nat_cpoType =-> (One:cpoType) + nat_cpoType :=
|
||||
Eval hnf in mk_fcont zeroCase_cont.
|
||||
(*=End *)
|
||||
|
||||
Lemma zeroCase_simplS: forall n, zeroCase (S n) = @in2 _ _ (discrete_cpoType nat) n.
|
||||
intros n ; auto.
|
||||
Qed.
|
||||
|
||||
Lemma zeroCase_simplO: zeroCase O = @in1 _ (One:cpoType) _ tt.
|
||||
auto.
|
||||
Qed.
|
||||
|
||||
Lemma SimpleB_mon (A B :Type) (C:ordType) (op : A -> B -> C:Type) : monotonic (fun p:discrete_ordType A * discrete_ordType B => op (fst p) (snd p)).
|
||||
case => x0 y0. case => x1 y1. case ; simpl. move => L L'.
|
||||
have e:x0 = x1 by []. have e':y0 = y1 by []. rewrite e. by rewrite e'.
|
||||
Qed.
|
||||
|
||||
Definition SimpleBOpm (A B :Type) (C:ordType) (op : A -> B -> C:Type) : discrete_ordType A * discrete_ordType B =-> C :=
|
||||
Eval hnf in mk_fmono (SimpleB_mon op).
|
||||
|
||||
Lemma SimpleB_cont (A B:Type) (C:cpoType) (op : A -> B -> C:Type) :
|
||||
@continuous (discrete_cpoType A * discrete_cpoType B) C (SimpleBOpm op).
|
||||
move => c. simpl. apply: (Ole_trans _ (le_lub _ 0)). simpl.
|
||||
by [].
|
||||
Qed.
|
||||
|
||||
Definition SimpleBOp (A B:Type) (C:cpoType) (op : A -> B -> C:Type) : discrete_cpoType A * discrete_cpoType B =-> C :=
|
||||
Eval hnf in mk_fcont (SimpleB_cont op).
|
||||
|
||||
(*=SemValExp *)
|
||||
Fixpoint SemVal E (v:Value E) : SemEnv E =-> VInf :=
|
||||
match v return SemEnv E =-> VInf with
|
||||
| INT i => in1 << const _ i
|
||||
| VAR m => SemVar m
|
||||
| LAMBDA e => in2 << exp_fun (kleisli (eta << RD.Roll) << SemExp e << Id >< RD.Unroll)
|
||||
end with SemExp E (e:Exp E) : SemEnv E =-> RD.VInf _BOT :=
|
||||
match e with
|
||||
| VAL v => eta << SemVal v
|
||||
| APP v1 v2 => kleisli (eta << RD.Unroll) << ev <<
|
||||
<| [| @const _ (exp_cppoType _ _) PBot , Id|] << SemVal v1, RD.Roll << SemVal v2 |>
|
||||
| LET e1 e2 => ev << <| exp_fun (KLEISLIR (SemExp e2)), SemExp e1 |>
|
||||
| OP op v0 v1 => kleisli (eta << in1 << SimpleBOp op) << uncurry (Smash _ _) <<
|
||||
<| [| eta, const _ PBot|] << SemVal v0, [| eta, const _ PBot|] << SemVal v1|>
|
||||
| IFZ v e1 e2 => ev <<
|
||||
[| [| exp_fun (SemExp e1 << pi2), exp_fun (SemExp e2 << pi2)|] << zeroCase ,
|
||||
@const _ (exp_cppoType _ _) PBot |] >< Id << <|SemVal v, Id|>
|
||||
end.
|
||||
(*=End *)
|
||||
|
||||
Lemma Operator2_strictL A B C (f:A * B =-> C _BOT) d : Operator2 f PBot d =-= PBot.
|
||||
apply: Ole_antisym ; last by apply: leastP.
|
||||
unlock Operator2. simpl. by do 2 rewrite kleisli_bot.
|
||||
Qed.
|
||||
|
||||
Lemma Operator2_strictR A B C (f:A * B =-> C _BOT) d : Operator2 f d PBot =-= PBot.
|
||||
apply: Ole_antisym ; last by apply: leastP.
|
||||
unlock Operator2. simpl.
|
||||
apply: (Ole_trans (proj2 (fmon_eq_elim (kleisli_comp2 _ _) d))).
|
||||
apply: DLless_cond. move => c X.
|
||||
case: (kleisliValVal X) => a [e P]. rewrite -> e. clear d e X.
|
||||
simpl in P. rewrite -> kleisli_bot in P.
|
||||
by case: (PBot_incon (proj2 P)).
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism E (e:Exp E) : (SemExp e)
|
||||
with signature (@tset_eq _ : SemEnv E -> SemEnv E -> Prop) ==> (@tset_eq _ : VInf _BOT -> VInf _BOT -> Prop)
|
||||
as SemExp_eq_compat.
|
||||
intros e0 e1 eeq. by apply (fmon_stable (SemExp e)).
|
||||
Qed.
|
||||
|
||||
Add Parametric Morphism E (v:Value E) : (SemVal v)
|
||||
with signature (@tset_eq _ : SemEnv E -> SemEnv E -> Prop) ==> (@tset_eq _ : VInf -> VInf -> Prop)
|
||||
as SemVal_eq_compat.
|
||||
intros e0 e1 eeq. by apply (fmon_stable (SemVal v)).
|
||||
Qed.
|
||||
|
||||
|
171
papers/domains-2012/uniisound.v
Executable file
171
papers/domains-2012/uniisound.v
Executable file
|
@ -0,0 +1,171 @@
|
|||
(**********************************************************************************
|
||||
* uniisound.v *
|
||||
* Formalizing Domains, Ultrametric Spaces and Semantics of Programming Languages *
|
||||
* Nick Benton, Lars Birkedal, Andrew Kennedy and Carsten Varming *
|
||||
* Jan 2012 *
|
||||
* Build with Coq 8.3pl2 plus SSREFLECT *
|
||||
**********************************************************************************)
|
||||
|
||||
(* soundness of semantics of unityped lambda calculus *)
|
||||
|
||||
Require Import PredomAll.
|
||||
Require Import ssrnat.
|
||||
Require Import uniisem uniiop Fin.
|
||||
|
||||
Set Implicit Arguments.
|
||||
Unset Strict Implicit.
|
||||
Import Prenex Implicits.
|
||||
|
||||
Section SEMMAP.
|
||||
|
||||
Variable P : Env -> Type.
|
||||
Variable ops : Map.Ops P.
|
||||
Variable Sem : forall E, P E -> SemEnv E =-> VInf.
|
||||
Variable SemVl : forall E (v : P E), Sem v =-= SemVal (Map.vl ops v).
|
||||
Variable SemVr : forall E, Sem (Map.vr ops (FinZ E)) =-= pi2.
|
||||
Variable SemWk : forall E (v : P E), Sem (Map.wk ops v) =-= Sem v << pi1.
|
||||
|
||||
Fixpoint SemMap E E' : Map.Map P E' E -> SemEnv E =-> SemEnv E' :=
|
||||
match E' with
|
||||
| O => fun m => terminal_morph (SemEnv E)
|
||||
| S _ => fun m => <| SemMap (tl m), Sem (hd m) |>
|
||||
end.
|
||||
|
||||
Lemma SemShift : forall E E' (m : Map.Map P E E'), SemMap (Map.shift ops m) =-= SemMap m << pi1.
|
||||
Proof. elim.
|
||||
- move => E' m. by apply: terminal_unique.
|
||||
- move => E IH E' m.
|
||||
rewrite -> (consEta m) at 1. rewrite Map.shiftCons. simpl. rewrite -> (IH E' (tl m)). rewrite prod_fun_compl. unfold Map.shift, hd. by rewrite SemWk.
|
||||
Qed.
|
||||
|
||||
Lemma SemLift : forall E E' (m : Map.Map P E E'), SemMap (Map.lift ops m) =-= SemMap m >< Id.
|
||||
Proof. move => E E' m. simpl. unfold tl, hd. simpl Map.lift. rewrite SemVr. rewrite SemShift.
|
||||
apply: prod_unique. rewrite prod_fun_fst. by rewrite prod_fun_fst.
|
||||
rewrite prod_fun_snd. rewrite prod_fun_snd. by rewrite comp_idL.
|
||||
Qed.
|
||||
|
||||
Lemma SemId : forall E, SemMap (@Map.id P ops E) =-= Id.
|
||||
Proof. elim.
|
||||
- simpl. by apply: terminal_unique.
|
||||
- move => env IH. rewrite Map.idAsCons. simpl. rewrite tlCons hdCons. rewrite SemShift. rewrite IH. rewrite comp_idL. rewrite SemVr.
|
||||
apply: prod_unique. rewrite prod_fun_fst. by rewrite comp_idR. rewrite prod_fun_snd. by rewrite comp_idR.
|
||||
Qed.
|
||||
|
||||
Lemma SemCommutesWithMap E :
|
||||
(forall (v : Value E) E' (r : Map.Map _ E E'), SemVal v << SemMap r =-= SemVal (Map.mapVal ops r v))
|
||||
/\ (forall (e : Exp E) E' (r : Map.Map _ E E'), SemExp e << SemMap r =-= SemExp (Map.mapExp ops r e)).
|
||||
Proof. move: E ; apply ExpValue_ind.
|
||||
(* VAR *)
|
||||
- move => n v n' s. SimplMap. simpl.
|
||||
induction v.
|
||||
+ rewrite (consEta s). simpl. rewrite prod_fun_snd hdCons. by rewrite SemVl.
|
||||
+ rewrite (consEta s). simpl. rewrite <- IHv. rewrite <- comp_assoc. by rewrite prod_fun_fst tlCons.
|
||||
(* INT *)
|
||||
- move => n body n' s. SimplMap. simpl. rewrite <- comp_assoc. by rewrite <- const_com.
|
||||
(* LAMBDA *)
|
||||
- move => n body IH n' s. SimplMap. simpl SemVal.
|
||||
rewrite <- (IH _ (Map.lift _ s)). rewrite SemLift. do 4 rewrite <- comp_assoc. apply: comp_eq_compat ; first by [].
|
||||
apply: exp_unique. rewrite <- (comp_assoc pi1 (SemMap s)). rewrite <- prod_map_prod_fun.
|
||||
rewrite (comp_assoc _ _ ev). rewrite exp_com. do 2 rewrite <- comp_assoc. rewrite prod_map_prod_fun. rewrite comp_idL.
|
||||
rewrite prod_fun_compl. rewrite <- comp_assoc. rewrite prod_fun_fst. rewrite prod_fun_snd. by do 2 rewrite comp_idL.
|
||||
(* VAL *)
|
||||
- move => n v IH n' s. SimplMap. simpl. rewrite <- comp_assoc. by rewrite IH.
|
||||
(* APP *)
|
||||
- move => n v IH v' IH' n' s. SimplMap. simpl. rewrite <- IH. rewrite <- IH'.
|
||||
rewrite <- comp_assoc. rewrite prod_fun_compl. by repeat (rewrite <- comp_assoc).
|
||||
(* LET *)
|
||||
- move => n e IH e' IH' n' s. SimplMap. simpl.
|
||||
rewrite <- comp_assoc. rewrite prod_fun_compl. rewrite IH. rewrite exp_comp. rewrite <- IH'. rewrite SemLift.
|
||||
rewrite <- (comp_idL (SemExp (Map.mapExp ops s e))). rewrite <- (comp_idR (exp_fun _)).
|
||||
rewrite <- (prod_map_prod_fun (exp_fun (KLEISLIR (SemExp e') << SemMap s >< Id))).
|
||||
rewrite comp_assoc. rewrite exp_com. rewrite <- comp_assoc. rewrite prod_map_prod_fun.
|
||||
rewrite comp_idR. rewrite comp_idL.
|
||||
rewrite <- (comp_idR (exp_fun _)). rewrite <- (comp_idL (SemExp (Map.mapExp ops s e))). rewrite <- prod_map_prod_fun.
|
||||
rewrite comp_assoc. rewrite exp_com. rewrite KLEISLIR_comp. by do 2 rewrite comp_idL.
|
||||
(* IFZ *)
|
||||
- move => n v IH e0 IH0 e1 IH1 n' s.
|
||||
SimplMap. simpl. repeat rewrite <- comp_assoc. repeat rewrite prod_fun_compl.
|
||||
repeat rewrite comp_idL. do 2 rewrite prod_fun_snd. rewrite <- comp_assoc. rewrite prod_fun_fst.
|
||||
rewrite IH. clear IH. rewrite <- comp_assoc. rewrite prod_fun_fst.
|
||||
apply: fmon_eq_intro => x. simpl. case: ((SemVal (Map.mapVal _ s v)) x) ; simpl.
|
||||
+ case.
|
||||
* simpl. do 2 rewrite SUM_fun_simplx. simpl. do 2 rewrite SUM_fun_simplx. apply: (fmon_eq_elim (IH0 _ s) x).
|
||||
* simpl => nn. do 2 rewrite SUM_fun_simplx. simpl. do 2 rewrite SUM_fun_simplx. apply (fmon_eq_elim (IH1 _ s) x).
|
||||
+ simpl. move => m. do 2 rewrite SUM_fun_simplx. by split ; apply: leastP.
|
||||
(* OP *)
|
||||
- move => n op v0 IH0 v1 IH1 v s.
|
||||
SimplMap. simpl.
|
||||
by (repeat rewrite <- comp_assoc; repeat rewrite prod_fun_compl; repeat rewrite <- comp_assoc; rewrite IH1; rewrite IH0).
|
||||
Qed.
|
||||
|
||||
End SEMMAP.
|
||||
|
||||
Definition SemRen := SemMap SemVar.
|
||||
Definition SemSub := SemMap SemVal.
|
||||
|
||||
Lemma SemCommutesWithRen E:
|
||||
(forall (v : Value E) E' (r : Ren E E'), SemVal v << SemRen r =-= SemVal (renVal r v))
|
||||
/\ (forall (e : Exp E) E' (r : Ren E E'), SemExp e << SemRen r =-= SemExp (renExp r e)).
|
||||
Proof. by apply SemCommutesWithMap. Qed.
|
||||
|
||||
Lemma SemShiftRen : forall E E' (r : Ren E E'), SemRen (shiftRen r) =-= SemRen r << pi1.
|
||||
Proof. by apply SemShift. Qed.
|
||||
|
||||
Lemma SemIdRen : forall E, SemRen (@idRen E) =-= Id.
|
||||
Proof. by apply SemId. Qed.
|
||||
|
||||
|
||||
(*=Substitution *)
|
||||
Lemma SemCommutesWithSub E:
|
||||
(forall (v : Value E) E' (s : Sub E E'), SemVal v << SemSub s =-= SemVal (subVal s v))
|
||||
/\ (forall (e : Exp E) E' (s : Sub E E'), SemExp e << SemSub s =-= SemExp (subExp s e)).
|
||||
(*=End *)
|
||||
Proof. move: E ; apply SemCommutesWithMap.
|
||||
+ by []. + by []. + intros. simpl. rewrite <- (proj1 (SemCommutesWithRen E)). rewrite SemShiftRen. rewrite SemIdRen. by rewrite comp_idL.
|
||||
Qed.
|
||||
|
||||
(*=Soundness *)
|
||||
Lemma Soundness e v : (e =>> v) -> SemExp e =-= eta << SemVal v.
|
||||
(*=End *)
|
||||
Proof. move => D. elim: e v / D.
|
||||
- by [].
|
||||
- move => e v v' D IH. simpl.
|
||||
repeat rewrite <- comp_assoc. rewrite <- (comp_idR ([|_,_|] << _)). rewrite <- (comp_idL (Roll << _)).
|
||||
rewrite <- prod_map_prod_fun.
|
||||
rewrite (comp_assoc _ _ ev). rewrite (comp_assoc _ in2). rewrite sum_fun_snd.
|
||||
rewrite (comp_idL (exp_fun _)). rewrite exp_com. repeat rewrite <- comp_assoc.
|
||||
rewrite comp_assoc. rewrite <- kleisli_comp2. rewrite <- comp_assoc.
|
||||
rewrite UR_id. rewrite comp_idR. rewrite kleisli_unit. rewrite comp_idL.
|
||||
rewrite prod_map_prod_fun. rewrite comp_idR. rewrite comp_assoc. rewrite UR_id.
|
||||
rewrite comp_idL.
|
||||
rewrite <- (proj2 (SemCommutesWithSub _) e _ ([v] %sub)) in IH.
|
||||
rewrite <- IH. simpl. rewrite (terminal_unique Id _). by [].
|
||||
- move => e0 v0 e1 v1 D0 IH0 D1 IH1.
|
||||
simpl. rewrite IH0. clear IH0 D0.
|
||||
rewrite <- (comp_idL (eta << _)). rewrite <- (comp_idR (exp_fun _)). rewrite <- prod_map_prod_fun.
|
||||
rewrite comp_assoc. rewrite exp_com. rewrite KLEISLIR_unit.
|
||||
rewrite <- (proj2 (SemCommutesWithSub 1%N) e1 _ ([ v0]%sub)) in IH1. simpl in IH1.
|
||||
rewrite <- IH1. rewrite (terminal_unique Id (terminal_morph One)). by [].
|
||||
- move => e0 e1 v0 _ IH. simpl. repeat rewrite <- comp_assoc. rewrite prod_map_prod_fun.
|
||||
rewrite comp_assoc. rewrite sum_fun_fst. rewrite <- IH. clear IH.
|
||||
rewrite <- comp_assoc.
|
||||
have X:(zeroCase << @const One (discrete_cpoType nat) 0%N) =-= in1 by apply: fmon_eq_intro ; case.
|
||||
rewrite X. clear X. rewrite sum_fun_fst. rewrite <- (comp_idR (exp_fun _)). rewrite <- prod_map_prod_fun.
|
||||
rewrite comp_assoc. rewrite exp_com. rewrite <- comp_assoc. rewrite prod_fun_snd. by rewrite comp_idR.
|
||||
- move => e0 e1 v0 n _ IH. simpl. repeat rewrite <- comp_assoc. rewrite prod_map_prod_fun.
|
||||
rewrite comp_assoc. rewrite sum_fun_fst. rewrite <- IH. clear IH.
|
||||
rewrite <- comp_assoc.
|
||||
have X:(zeroCase << @const One (discrete_cpoType nat) n.+1%N) =-= in2 << @const One (discrete_cpoType nat) n by apply: fmon_eq_intro ; case.
|
||||
rewrite X. clear X. rewrite comp_assoc. rewrite sum_fun_snd.
|
||||
rewrite <- prod_map_prod_fun. rewrite comp_assoc. rewrite exp_com. rewrite <- comp_assoc. rewrite prod_fun_snd.
|
||||
by rewrite comp_idR.
|
||||
- move => op n n'. simpl. repeat rewrite comp_assoc. rewrite sum_fun_fst. simpl.
|
||||
rewrite <- (prod_map_prod_fun eta). do 2 rewrite <- comp_assoc.
|
||||
rewrite (comp_assoc _ _ ev). rewrite - {1} (comp_idL pi2).
|
||||
rewrite (comp_assoc _ (eta >< eta)).
|
||||
have e:((ev << Smash (discrete_cpoType nat) (discrete_cpoType nat) >< Id) <<
|
||||
eta >< eta) =-= eta by apply: fmon_eq_intro => a; simpl ; unfold Smash ; simpl ; rewrite Operator2_simpl ; case: a.
|
||||
rewrite -> e. clear e. rewrite comp_assoc. rewrite kleisli_eta_com.
|
||||
rewrite <- comp_assoc. by apply:fmon_eq_intro.
|
||||
Qed.
|
||||
|
|
@ -34,6 +34,7 @@ open Eq.≡-Reasoning using (begin_; _≡⟨⟩_; _≡⟨_⟩_; _∎)
|
|||
open import Data.Nat using (ℕ; zero; suc; _+_; _*_; _∸_; _≤_; z≤n; s≤s)
|
||||
open import Data.Nat.Properties using (+-assoc; +-identityʳ; +-suc; +-comm;
|
||||
≤-refl; ≤-trans; ≤-antisym; ≤-total; +-monoʳ-≤; +-monoˡ-≤; +-mono-≤)
|
||||
open import plfa.Relations using (_<_; z<s; s<s)
|
||||
\end{code}
|
||||
|
||||
## Naturals
|
||||
|
@ -144,6 +145,11 @@ the following function from the standard library:
|
|||
|
||||
sym : ∀ {m n : ℕ} → m ≡ n → n ≡ m
|
||||
|
||||
\begin{code}
|
||||
swap : ∀ (m n p : ℕ) → m + (n + p) ≡ n + (m + p)
|
||||
swap m n p rewrite sym (+-assoc m n p) | +-comm m n | +-assoc n m p = {!!}
|
||||
\end{code}
|
||||
|
||||
#### Exercise `*-distrib-+` (recommended) {#times-distrib-plus}
|
||||
|
||||
Show multiplication distributes over addition, that is,
|
||||
|
|
|
@ -101,6 +101,8 @@ _ = s≤s (s≤s z≤n)
|
|||
\end{code}
|
||||
|
||||
|
||||
|
||||
|
||||
## Implicit arguments
|
||||
|
||||
This is our first use of implicit arguments. In the definition of
|
||||
|
|
Loading…
Add table
Reference in a new issue