@@ -1637,6 +1637,68 @@ Qed.
1637
1637
G |-- Rec (F P) >=> Rec (F Q).
1638
1638
*)
1639
1639
1640
+ Section HORec_sub_strong.
1641
+ Variable A B : Type.
1642
+ Variable NA : NatDed A.
1643
+ Variable IA : Indir A.
1644
+ Variable RA : RecIndir A.
1645
+ Variable F:(B -> A) -> (B -> A) -> B -> A.
1646
+ Variable HF1 : forall (f:B->A), HOcontractive (F f).
1647
+ Variable HF2 : forall (R: B -> A) (P Q: B->A), ALL b, P b >=> Q b |-- ALL b, F P R b >=> F Q R b.
1648
+ Variable HF3 : forall (P Q: B -> A) (f:B->A), ALL b:B, |>(P b >=> Q b) |-- ALL b:B, F f P b >=> F f Q b.
1649
+
1650
+ Lemma HORec_sub_strong G (f g : B->A) (H: G |-- ALL b:B, (f b) >=> (g b)):
1651
+ G |-- ALL b:B, HORec (F f) b >=> HORec (F g) b.
1652
+ Proof .
1653
+ assert (HF2': forall (R: B -> A) b (P Q: B->A), ALL a, P a >=> Q a |-- F P R b >=> F Q R b).
1654
+ { intros. eapply derives_trans. apply (HF2 R P Q).
1655
+ apply allp_left with b. trivial. }
1656
+ clear HF2.
1657
+ apply @derives_trans with (ALL b, f b >=> g b); auto.
1658
+ clear G H.
1659
+ apply goedel_loeb.
1660
+ apply allp_right; intro b.
1661
+ rewrite HORec_fold_unfold by auto.
1662
+ pose proof (HORec_fold_unfold _ _ (HF1 f)).
1663
+ pose proof (HORec_fold_unfold _ _ (HF1 g)).
1664
+ set (P' := HORec (F f)) in *.
1665
+ set (Q' := HORec (F g)) in *.
1666
+ rewrite <- H.
1667
+ specialize (HF3 P' Q' f).
1668
+ rewrite later_allp.
1669
+ eapply derives_trans; [apply andp_derives ; [apply derives_refl | apply HF3] | ].
1670
+ specialize (HF2' Q' b f g). rewrite <- H0 in HF2'.
1671
+ rewrite <- H in *.
1672
+ apply subp_trans with (F f Q' b).
1673
+ apply andp_left2. apply allp_left with b; auto.
1674
+ apply andp_left1; auto.
1675
+ Qed .
1676
+
1677
+ End HORec_sub_strong.
1678
+
1679
+ Lemma HORec_sub' {A} {NA: NatDed A}{IA: Indir A}{RA: RecIndir A} : forall G B
1680
+ (F : A -> (B -> A) -> B -> A)
1681
+ (HF1 : forall (X:A), HOcontractive (F X))
1682
+ (HF2 : forall (R: B -> A) (P Q: A), P >=> Q |-- ALL b, F P R b >=> F Q R b)
1683
+ (HF3 : forall (P Q: B -> A) (X:A), ALL b:B, |>(P b >=> Q b) |-- ALL b:B, F X P b >=> F X Q b),
1684
+ forall P Q : A,
1685
+ (G |-- P >=> Q) ->
1686
+ G |-- ALL b:B, HORec (F P) b >=> HORec (F Q) b.
1687
+ Proof . intros.
1688
+ apply (@HORec_sub_strong A B NA IA RA (fun f g b => F (f b) g b)).
1689
+ + clear - HF1. red; intros.
1690
+ apply allp_right; intros b.
1691
+ eapply derives_trans; [ apply (HF1 (f b) P Q) | clear HF1].
1692
+ apply allp_left with b; trivial.
1693
+ + clear - HF2. intros R f g. apply allp_derives; intros b.
1694
+ eapply derives_trans; [ apply (HF2 R (f b) (g b)) | clear HF2].
1695
+ apply allp_left with b; trivial.
1696
+ + clear - HF3. intros P Q f. apply allp_right; intros b.
1697
+ eapply derives_trans; [ apply (HF3 P Q (f b)) | clear HF3].
1698
+ apply allp_left with b; trivial.
1699
+ + apply allp_right; intros b; trivial.
1700
+ Qed .
1701
+
1640
1702
Lemma HORec_sub {A} {NA: NatDed A}{IA: Indir A}{RA: RecIndir A} : forall G B
1641
1703
(F : A -> (B -> A) -> B -> A)
1642
1704
(HF1 : forall X, HOcontractive (F X))
@@ -1646,6 +1708,9 @@ Lemma HORec_sub {A} {NA: NatDed A}{IA: Indir A}{RA: RecIndir A} : forall G B
1646
1708
(G |-- P >=> Q) ->
1647
1709
G |-- ALL b:B, HORec (F P) b >=> HORec (F Q) b.
1648
1710
Proof .
1711
+ intros. apply HORec_sub'; trivial.
1712
+ intros. apply allp_right; intros b; apply HF2.
1713
+ (*old proof:
1649
1714
intros.
1650
1715
apply @derives_trans with (P>=>Q); auto.
1651
1716
clear G H.
@@ -1664,10 +1729,10 @@ Proof.
1664
1729
rewrite <- H in *.
1665
1730
apply subp_trans with (F P Q' b).
1666
1731
apply andp_left2. apply allp_left with b; auto.
1667
- apply andp_left1; auto.
1732
+ apply andp_left1. auto.
1733
+ *)
1668
1734
Qed .
1669
1735
1670
-
1671
1736
(****** End contractiveness **** *)
1672
1737
1673
1738
Require Import Coq.ZArith.ZArith.
0 commit comments