From c6af4e5d740de9aabe0d96780ae77a907d77a259 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 27 Jan 2023 16:54:37 +0900 Subject: [PATCH 01/82] add MonadTypedStore --- hierarchy.v | 54 +++++++++++++++++++ monad_model.v | 141 ++++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 195 insertions(+) diff --git a/hierarchy.v b/hierarchy.v index 0edb9016..62c0eab9 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1076,6 +1076,60 @@ HB.structure Definition MonadArray (S : UU0) (I : eqType) := HB.structure Definition MonadPlusArray (S : UU0) (I : eqType) := { M of MonadPlus M & isMonadArray S I M}. +(* A monad for OCaml computations *) +Module Type MLTY. +Parameter ml_type : Set. +Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. +Variant loc : ml_type -> Type := mkloc T : nat -> loc T. +Parameter coq_type : forall M : Type -> Type, ml_type -> Type. +End MLTY. + +Module MonadTypedStore (MLtypes : MLTY). +Import MLtypes. + +Record binding (M : Type -> Type) := + mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. +Arguments mkbind {M}. + +Definition loc_id {T} (r : loc T) : nat := let: mkloc _ n := r in n. + +HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) + of Monad M := { + cnew : forall {T}, coq_type M T -> M (loc T) ; + cget : forall {T}, loc T -> M (coq_type M T) ; + cput : forall {T}, loc T -> coq_type M T -> M unit ; + cchk : forall {T}, loc T -> M unit ; + cputput : forall T (r : loc T) (s s' : coq_type M T), + cput r s >> cput r s' = cput r s' ; + cputget : + forall T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A), + cput r s >> cget r >>= k = cput r s >> k s ; + cgetputchk : forall T (r : loc T), cget r >>= cput r = cchk r ; + cgetget : + forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), + cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s ; + cgetC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) + (k : coq_type M T1 -> coq_type M T2 -> M A), + cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = + cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)) ; + cputC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (s2 : coq_type M T2) (A : UU0), + loc_id r1 != loc_id r2 \/ mkbind T1 s1 = mkbind T2 s2 -> + cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1 ; + cputgetC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (A : UU0) (k : coq_type M T2 -> M A), + loc_id r1 != loc_id r2 -> + cput r1 s1 >> cget r2 >>= k = + cget r2 >>= (fun v => cput r1 s1 >> k v) }. + +#[short(type=typedStoreMonad)] +HB.structure Definition MonadTypedStore := + { M of isMonadTypedStore M & isMonad M & isFunctor M }. +End MonadTypedStore. + HB.mixin Record isMonadTrace (T : UU0) (M : UU0 -> UU0) of Monad M := { mark : T -> M unit }. diff --git a/monad_model.v b/monad_model.v index fc73811c..ba7cf508 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1656,6 +1656,147 @@ End modelplusarray. End ModelPlusArray. HB.export ModelPlusArray. +Module ModelTypedStore (MLtypes : MLTY). +Module MTypedStore := MonadTypedStore (MLtypes). +Import MLtypes. +Import MTypedStore. + +#[bypass_check(positivity)] +Inductive acto : UU0 -> UU0 := + mkActo : forall T : UU0, MS (seq (binding acto)) option_monad T -> acto T. +Local Notation M := acto. +Local Notation coq_type := (coq_type M). + +Definition ofActo T (m : M T) : MS (seq (binding M)) option_monad T := + let: mkActo _ m' := m in m'. + +Definition cnew T (v : coq_type T) : M (loc T) := + mkActo (fun st => let n := size st in Ret (mkloc T n, rcons st (mkbind T v))). + +Definition coerce (T1 T2 : ml_type) (v : coq_type T1) : option (coq_type T2) := + match ml_type_eq_dec T1 T2 with + | left H => Some (eq_rect _ _ v _ H) + | right _ => None + end. + +Definition cget T (r : loc T) : M (coq_type T) := + mkActo (fun st => + if List.nth_error st (loc_id r) is Some (mkbind T' v) then + if coerce T v is Some u then Ret (u, st) else fail + else fail). + +Definition cput T (r : loc T) (v : coq_type T) : M unit := + mkActo (fun st => + let n := loc_id r in + if List.nth_error st n is Some (mkbind T' _) then + if coerce T' v is Some u then + let b := mkbind T' u in + Ret (tt, set_nth b st n b) + else fail + else fail). + +Definition cchk T (r : loc T) : M unit := + mkActo (fun st => + if List.nth_error st (loc_id r) is Some (mkbind T' u) then + if coerce T u is Some _ then Ret (tt, st) else fail + else fail). + +Let ret : forall A, idfun A -> M A := fun A a => mkActo (Ret a). +Let bind A B (m : M A) (f : A -> M B) : M B := + mkActo (ofActo m >>= (fun a => ofActo (f a))). +Let left_neutral : BindLaws.left_neutral bind ret. +Proof. by move=> A B a f; rewrite /bind /ret bindretf; case: (f a). Qed. +Let right_neutral : BindLaws.right_neutral bind ret. +Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. +Let associative : BindLaws.associative bind. +Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. + +(* +HB.instance Definition xyz := + isMonad_ret_bind.Build M left_neutral right_neutral associative. +*) + +Lemma nth_error_set_nth T (x : T) st n : + List.nth_error (set_nth x st n x) n = Some x. +Proof. +elim: n st => [|z IH] [] //. +clear IH. +elim: z.+1 => [|n <-] //=. +by rewrite set_nth_nil. +Qed. + +Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. +Lemma ml_type_eqP : Equality.axiom ml_type_eqb. +Proof. +rewrite /ml_type_eqb => T1 T2. +by case: ml_type_eq_dec; constructor. +Qed. +Definition ml_type_eq_mixin := EqMixin ml_type_eqP. +Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. + +Let cputput T (r : loc T) (s s' : coq_type T) : + bind (cput r s) (fun _ => cput r s') = cput r s'. +Proof. +congr mkActo. +apply/boolp.funext => st. +case: r s s' => {}T n s s' /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hst : (List.nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T'. +rewrite !bindretf /= nth_error_set_nth /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +set b := mkbind _ _. +set b' := mkbind _ _. +admit. +Admitted. + +Let cputget T (r : loc T) (s : coq_type T) (A : UU0) (k : coq_type T -> M A) : + bind (cput r s) (fun=> bind (cget r) k) = bind (cput r s) (fun=> k s). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +case: r s k => {}T n s k /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hst : (List.nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T'. +rewrite !bindretf /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite nth_error_set_nth. +rewrite /coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +by rewrite -eq_rect_eq !bindretf. +Qed. + +(* + cgetputchk : forall T (r : loc T), cget r >>= cput r = cchk r ; + cgetget : + forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), + cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s ; + cgetC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) + (k : coq_type M T1 -> coq_type M T2 -> M A), + cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = + cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)) ; + cputC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (s2 : coq_type M T2) (A : UU0), + loc_id r1 != loc_id r2 \/ mkbind T1 s1 = mkbind T2 s2 -> + cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1 ; + cputgetC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (A : UU0) (k : coq_type M T2 -> M A), + loc_id r1 != loc_id r2 -> + cput r1 s1 >> cget r2 >>= k = + cget r2 >>= (fun v => cput r1 s1 >> k v) }. +*) +End ModelTypedStore. + (* TODO? (* result of a discussion with Maxime and Enrico on 2019-09-12 *) (* Equality between monads from the hierarchy and their counterparts built *) From acea7fe2ff91c65a1eadadc1926dceea8e5c98b5 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 27 Jan 2023 17:04:07 +0900 Subject: [PATCH 02/82] add axioms for new and get --- hierarchy.v | 10 +++++++++- 1 file changed, 9 insertions(+), 1 deletion(-) diff --git a/hierarchy.v b/hierarchy.v index 62c0eab9..a9b3cc59 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1123,7 +1123,15 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) (A : UU0) (k : coq_type M T2 -> M A), loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = - cget r2 >>= (fun v => cput r1 s1 >> k v) }. + cget r2 >>= (fun v => cput r1 s1 >> k v) ; + cnewget : forall T (s : coq_type M T) A (k : coq_type M T -> M A), + cnew s >>= (fun r => cget r >>= k) = cnew s >> k s ; + cnewgetC : + forall T T' (r : loc T) (s : coq_type M T') A + (k : loc T' -> coq_type M T -> M A), + cnew s >>= (fun r' => cget r >>= k r') = + cget r >>= (fun u => cnew s >>= (fun r' => k r' u)) ; + }. #[short(type=typedStoreMonad)] HB.structure Definition MonadTypedStore := From 4b12c3fce8c2abcd905e7e6f72227bf92632ecfd Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Fri, 27 Jan 2023 17:17:34 +0900 Subject: [PATCH 03/82] sync with impredicative_set --- impredicative_set/ihierarchy.v | 62 ++++++++++++++ impredicative_set/imonad_model.v | 139 +++++++++++++++++++++++++++++++ monad_model.v | 4 +- 3 files changed, 203 insertions(+), 2 deletions(-) diff --git a/impredicative_set/ihierarchy.v b/impredicative_set/ihierarchy.v index c881f9a4..8c3b99b4 100644 --- a/impredicative_set/ihierarchy.v +++ b/impredicative_set/ihierarchy.v @@ -1065,6 +1065,68 @@ HB.structure Definition MonadArray (S : UU0) (I : eqType) := HB.structure Definition MonadPlusArray (S : UU0) (I : eqType) := { M of MonadPlus M & isMonadArray S I M}. +(* A monad for OCaml computations *) +Module Type MLTY. +Parameter ml_type : UU0. +Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. +Variant loc : ml_type -> UU0 := mkloc T : nat -> loc T. +Parameter coq_type : forall M : UU0 -> UU0, ml_type -> UU0. +End MLTY. + +Module MonadTypedStore (MLtypes : MLTY). +Import MLtypes. + +Record binding (M : UU0 -> UU0) := + mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. +Arguments mkbind {M}. + +Definition loc_id {T} (r : loc T) : nat := let: mkloc _ n := r in n. + +HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) + of Monad M := { + cnew : forall {T}, coq_type M T -> M (loc T) ; + cget : forall {T}, loc T -> M (coq_type M T) ; + cput : forall {T}, loc T -> coq_type M T -> M unit ; + cchk : forall {T}, loc T -> M unit ; + cputput : forall T (r : loc T) (s s' : coq_type M T), + cput r s >> cput r s' = cput r s' ; + cputget : + forall T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A), + cput r s >> cget r >>= k = cput r s >> k s ; + cgetputchk : forall T (r : loc T), cget r >>= cput r = cchk r ; + cgetget : + forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), + cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s ; + cgetC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) + (k : coq_type M T1 -> coq_type M T2 -> M A), + cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = + cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)) ; + cputC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (s2 : coq_type M T2) (A : UU0), + loc_id r1 != loc_id r2 \/ mkbind T1 s1 = mkbind T2 s2 -> + cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1 ; + cputgetC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (A : UU0) (k : coq_type M T2 -> M A), + loc_id r1 != loc_id r2 -> + cput r1 s1 >> cget r2 >>= k = + cget r2 >>= (fun v => cput r1 s1 >> k v) ; + cnewget : forall T (s : coq_type M T) A (k : coq_type M T -> M A), + cnew s >>= (fun r => cget r >>= k) = cnew s >> k s ; + cnewgetC : + forall T T' (r : loc T) (s : coq_type M T') A + (k : loc T' -> coq_type M T -> M A), + cnew s >>= (fun r' => cget r >>= k r') = + cget r >>= (fun u => cnew s >>= (fun r' => k r' u)) ; + }. + +#[short(type=typedStoreMonad)] +HB.structure Definition MonadTypedStore := + { M of isMonadTypedStore M & isMonad M & isFunctor M }. +End MonadTypedStore. + HB.mixin Record isMonadTrace (T : UU0) (M : UU0 -> UU0) of Monad M := { mark : T -> M unit }. diff --git a/impredicative_set/imonad_model.v b/impredicative_set/imonad_model.v index 24659216..a422caae 100644 --- a/impredicative_set/imonad_model.v +++ b/impredicative_set/imonad_model.v @@ -1118,6 +1118,145 @@ HB.instance Definition _ := isMonadArray.Build End modelarray. End ModelArray. +Module ModelTypedStore (MLtypes : MLTY). +Module MTypedStore := MonadTypedStore (MLtypes). +Import MLtypes. +Import MTypedStore. + +#[bypass_check(positivity)] +Inductive acto : UU0 -> UU0 := + mkActo : forall T : UU0, MS (seq (binding acto)) [the monad of option_monad] T -> acto T. +Local Notation M := acto. +Local Notation coq_type := (coq_type M). + +Definition ofActo (T:UU0) (m : M T) : MS (seq (binding M)) [the monad of option_monad] T := + let: mkActo _ m' := m in m'. + +Definition cnew T (v : coq_type T) : M (loc T) := + mkActo (fun st => let n := size st in Ret (mkloc T n, rcons st (mkbind T v))). + +Definition coerce (T1 T2 : ml_type) (v : coq_type T1) : option (coq_type T2) := + match ml_type_eq_dec T1 T2 with + | left H => Some (eq_rect _ _ v _ H) + | right _ => None + end. + +Definition cget T (r : loc T) : M (coq_type T) := + mkActo (fun st => + if List.nth_error st (loc_id r) is Some (mkbind T' v) then + if coerce T v is Some u then Ret (u, st) else fail + else fail). + +Definition cput T (r : loc T) (v : coq_type T) : M unit := + mkActo (fun st => + let n := loc_id r in + if List.nth_error st n is Some (mkbind T' _) then + if coerce T' v is Some u then + let b := mkbind T' u in + Ret (tt, set_nth b st n b) + else fail + else fail). + +Definition cchk T (r : loc T) : M unit := + mkActo (fun st => + if List.nth_error st (loc_id r) is Some (mkbind T' u) then + if coerce T u is Some _ then Ret (tt, st) else fail + else fail). + +Let ret : forall A : UU0, idfun A -> M A := fun A a => mkActo (Ret a). +Let bind (A B : UU0) (m : M A) (f : A -> M B) : M B := + mkActo (ofActo m >>= (fun a => ofActo (f a))). +Let left_neutral : BindLaws.left_neutral bind ret. +Proof. by move=> A B a f; rewrite /bind /ret bindretf; case: (f a). Qed. +Let right_neutral : BindLaws.right_neutral bind ret. +Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. +Let associative : BindLaws.associative bind. +Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. + +HB.instance Definition xyz := + isMonad_ret_bind.Build M left_neutral right_neutral associative. + +Lemma nth_error_set_nth T (x : T) st n : + List.nth_error (set_nth x st n x) n = Some x. +Proof. +elim: n st => [|z IH] [] //. +clear IH. +elim: z.+1 => [|n <-] //=. +by rewrite set_nth_nil. +Qed. + +Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. +Lemma ml_type_eqP : Equality.axiom ml_type_eqb. +Proof. +rewrite /ml_type_eqb => T1 T2. +by case: ml_type_eq_dec; constructor. +Qed. +Definition ml_type_eq_mixin := EqMixin ml_type_eqP. +Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. + +Let cputput T (r : loc T) (s s' : coq_type T) : + bind (cput r s) (fun _ => cput r s') = cput r s'. +Proof. +congr mkActo. +apply/funext => st. +case: r s s' => {}T n s s' /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hst : (List.nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T'. +rewrite !bindretf /= nth_error_set_nth /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +set b := mkbind _ _. +set b' := mkbind _ _. +admit. +Admitted. + +Let cputget T (r : loc T) (s : coq_type T) (A : UU0) (k : coq_type T -> M A) : + bind (cput r s) (fun=> bind (cget r) k) = bind (cput r s) (fun=> k s). +Proof. +congr mkActo. +apply/funext => st /=. +case: r s k => {}T n s k /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hst : (List.nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T'. +rewrite !bindretf /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite nth_error_set_nth. +rewrite /coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +by rewrite -eq_rect_eq !bindretf. +Qed. + +(* + cgetputchk : forall T (r : loc T), cget r >>= cput r = cchk r ; + cgetget : + forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), + cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s ; + cgetC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) + (k : coq_type M T1 -> coq_type M T2 -> M A), + cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = + cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)) ; + cputC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (s2 : coq_type M T2) (A : UU0), + loc_id r1 != loc_id r2 \/ mkbind T1 s1 = mkbind T2 s2 -> + cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1 ; + cputgetC : + forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (A : UU0) (k : coq_type M T2 -> M A), + loc_id r1 != loc_id r2 -> + cput r1 s1 >> cget r2 >>= k = + cget r2 >>= (fun v => cput r1 s1 >> k v) }. +*) +End ModelTypedStore. + (* TODO? (* result of a discussion with Maxime and Enrico on 2019-09-12 *) (* Equality between monads from the hierarchy and their counterparts built *) diff --git a/monad_model.v b/monad_model.v index ba7cf508..87962db5 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1663,11 +1663,11 @@ Import MTypedStore. #[bypass_check(positivity)] Inductive acto : UU0 -> UU0 := - mkActo : forall T : UU0, MS (seq (binding acto)) option_monad T -> acto T. + mkActo : forall T : UU0, MS (seq (binding acto)) [the monad of option_monad] T -> acto T. Local Notation M := acto. Local Notation coq_type := (coq_type M). -Definition ofActo T (m : M T) : MS (seq (binding M)) option_monad T := +Definition ofActo T (m : M T) : MS (seq (binding M)) [the monad of option_monad] T := let: mkActo _ m' := m in m'. Definition cnew T (v : coq_type T) : M (loc T) := From 7334c274f5768fb8fffa363b8f92cd37e648f4ae Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Sat, 28 Jan 2023 17:55:55 +0900 Subject: [PATCH 04/82] add cnew and cchk laws --- hierarchy.v | 41 +++- monad_model.v | 503 ++++++++++++++++++++++++++++++++++++++++++++++---- 2 files changed, 494 insertions(+), 50 deletions(-) diff --git a/hierarchy.v b/hierarchy.v index a9b3cc59..c21682d5 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -2,7 +2,7 @@ (* Copyright (C) 2020 monae authors, license: LGPL-2.1-or-later *) Ltac typeof X := type of X. -Require Import ssrmatching Reals. +Require Import ssrmatching Reals JMeq. From mathcomp Require Import all_ssreflect. From mathcomp Require boolp. From infotheo Require Import Reals_ext. @@ -1099,6 +1099,15 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cget : forall {T}, loc T -> M (coq_type M T) ; cput : forall {T}, loc T -> coq_type M T -> M unit ; cchk : forall {T}, loc T -> M unit ; + cnewget : forall T (s : coq_type M T) A (k : coq_type M T -> M A), + cnew s >>= (fun r => cget r >>= k) = cnew s >> k s ; + cnewgetC : + forall T T' (r : loc T) (s : coq_type M T') A + (k : loc T' -> coq_type M T -> M A), + cchk r >> (cnew s >>= (fun r' => cget r >>= k r')) = + cget r >>= (fun u => cnew s >>= (fun r' => k r' u)) ; + cnewput : forall T (s t : coq_type M T) A (k : loc T -> coq_type M T -> M A), + cnew s >>= (fun r => cput r t >> k r t) = cnew t >>= (fun r => k r t) ; cputput : forall T (r : loc T) (s s' : coq_type M T), cput r s >> cput r s' = cput r s' ; cputget : @@ -1116,7 +1125,7 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cputC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) (s2 : coq_type M T2) (A : UU0), - loc_id r1 != loc_id r2 \/ mkbind T1 s1 = mkbind T2 s2 -> + loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1 ; cputgetC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) @@ -1124,13 +1133,27 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v) ; - cnewget : forall T (s : coq_type M T) A (k : coq_type M T -> M A), - cnew s >>= (fun r => cget r >>= k) = cnew s >> k s ; - cnewgetC : - forall T T' (r : loc T) (s : coq_type M T') A - (k : loc T' -> coq_type M T -> M A), - cnew s >>= (fun r' => cget r >>= k r') = - cget r >>= (fun u => cnew s >>= (fun r' => k r' u)) ; + cnewchk : + forall T (s : coq_type M T) (A : UU0) (k : coq_type M T -> loc T -> M A), + cnew s >>= (fun r => cchk r >> k s r) = cnew s >>= k s ; + cchknewC : + forall T1 T2 (r : loc T1) (s : coq_type M T2) (A : UU0) + (k : coq_type M T2 -> loc T2 -> M A), + cchk r >> (cnew s >>= fun r' => cchk r >> k s r') = + cchk r >> (cnew s >>= k s) ; + cchkgetC : + forall T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type M T2 -> M A), + cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s) ; + cchkget : forall T (r : loc T) (A: UU0) (k : coq_type M T -> M A), + cchk r >> (cget r >>= k) = cget r >>= k ; + cgetchk : forall T (r : loc T) (A: UU0) (k : coq_type M T -> M A), + cget r >>= (fun s => cchk r >> k s) = cget r >>= k ; + cchkputC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2), + cchk r1 >> cput r2 s = cput r2 s >> cchk r1 ; + cchkput : forall T (r : loc T) (s : coq_type M T), + cchk r >> cput r s = cput r s ; + cputchk : forall T (r : loc T) (s : coq_type M T), + cput r s >> cchk r = cput r s ; }. #[short(type=typedStoreMonad)] diff --git a/monad_model.v b/monad_model.v index 87962db5..b0f841aa 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1656,20 +1656,33 @@ End modelplusarray. End ModelPlusArray. HB.export ModelPlusArray. -Module ModelTypedStore (MLtypes : MLTY). +Require Import JMeq. + +Module Type MLTYdef. +Include MLTY. +Parameter ml_undef : ml_type. +Parameter undef : forall M, coq_type M ml_undef. +End MLTYdef. + +Module ModelTypedStore (MLtypes : MLTYdef). Module MTypedStore := MonadTypedStore (MLtypes). Import MLtypes. Import MTypedStore. #[bypass_check(positivity)] Inductive acto : UU0 -> UU0 := - mkActo : forall T : UU0, MS (seq (binding acto)) [the monad of option_monad] T -> acto T. + mkActo : forall T : UU0, + MS (seq (binding acto)) [the monad of option_monad] T -> acto T. Local Notation M := acto. Local Notation coq_type := (coq_type M). -Definition ofActo T (m : M T) : MS (seq (binding M)) [the monad of option_monad] T := +Definition ofActo T (m : M T) + : MS (seq (binding M)) [the monad of option_monad] T := let: mkActo _ m' := m in m'. +Definition def := mkbind ml_undef (undef M). +Local Notation nth_error := List.nth_error. + Definition cnew T (v : coq_type T) : M (loc T) := mkActo (fun st => let n := size st in Ret (mkloc T n, rcons st (mkbind T v))). @@ -1681,23 +1694,22 @@ Definition coerce (T1 T2 : ml_type) (v : coq_type T1) : option (coq_type T2) := Definition cget T (r : loc T) : M (coq_type T) := mkActo (fun st => - if List.nth_error st (loc_id r) is Some (mkbind T' v) then + if nth_error st (loc_id r) is Some (mkbind T' v) then if coerce T v is Some u then Ret (u, st) else fail else fail). Definition cput T (r : loc T) (v : coq_type T) : M unit := mkActo (fun st => let n := loc_id r in - if List.nth_error st n is Some (mkbind T' _) then + if nth_error st n is Some (mkbind T' _) then if coerce T' v is Some u then - let b := mkbind T' u in - Ret (tt, set_nth b st n b) + Ret (tt, set_nth def st n (mkbind T' u)) else fail else fail). Definition cchk T (r : loc T) : M unit := mkActo (fun st => - if List.nth_error st (loc_id r) is Some (mkbind T' u) then + if nth_error st (loc_id r) is Some (mkbind T' u) then if coerce T u is Some _ then Ret (tt, st) else fail else fail). @@ -1711,13 +1723,17 @@ Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. Let associative : BindLaws.associative bind. Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. -(* +(* Doesn't work HB.instance Definition xyz := isMonad_ret_bind.Build M left_neutral right_neutral associative. *) +(* Since we couldn't build the instance, redefine notations *) +Local Notation "m >>= f" := (bind m f). +Local Notation "m >> f" := (bind m (fun=> f)). -Lemma nth_error_set_nth T (x : T) st n : - List.nth_error (set_nth x st n x) n = Some x. +(* Basic lemmas for standard library's nth_error *) +Lemma nth_error_set_nth T (def x : T) st n : + nth_error (set_nth def st n x) n = Some x. Proof. elim: n st => [|z IH] [] //. clear IH. @@ -1725,6 +1741,36 @@ elim: z.+1 => [|n <-] //=. by rewrite set_nth_nil. Qed. +Lemma nth_error_rcons_size T (st : seq T) b : + nth_error (rcons st b) (size st) = Some b. +Proof. by elim: st. Qed. + +Lemma nth_error_rcons_some T (st : seq T) n a b : + nth_error st n = Some a -> nth_error (rcons st b) n = Some a. +Proof. by elim: n st => [|n IH] []. Qed. + +Lemma nth_error_set_nth_id A (def : A) st n a : + nth_error st n = Some a -> set_nth def st n a = st. +Proof. elim: n st =>[|n IH] [] //= b st; by [case=> -> | move/IH ->]. Qed. + +Lemma nth_error_set_nth_other A (def : A) st m n a b : + m != n -> + nth_error st m = Some a -> + nth_error (set_nth def st n b) m = Some a. +Proof. elim: m st n => [|m IH] [|c st] [|n] //=; rewrite eqSS => *; exact: IH. +Qed. + +Lemma nth_error_set_nth_none A (def : A) st m n a b : + nth_error st m = None -> + nth_error st n = Some a -> + nth_error (set_nth def st n b) m = None. +Proof. by elim: m st n => [|m IH] [|c st] [|n] //=; apply IH. Qed. + +Lemma set_nth_rcons A (def : A) st a b : + set_nth def (rcons st a) (size st) b = rcons st b. +Proof. by elim: st => //= c st ->. Qed. + +(* Make ml_type an eqType *) Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. Lemma ml_type_eqP : Equality.axiom ml_type_eqb. Proof. @@ -1734,34 +1780,248 @@ Qed. Definition ml_type_eq_mixin := EqMixin ml_type_eqP. Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. +(* Prove the laws *) +Let cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : + cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite nth_error_rcons_size. +rewrite /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cnewgetC T T' (r : loc T) (s : coq_type T') A + (k : loc T' -> coq_type T -> M A) : + cchk r >> (cnew s >>= (fun r' => cget r >>= k r')) = + cget r >>= (fun u => cnew s >>= fun r' => k r' u). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T1. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite (nth_error_rcons_some _ Hr). +rewrite /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cnewput T (s t : coq_type T) A (k : loc T -> coq_type T -> M A) : + cnew s >>= (fun r => cput r t >> k r t) = cnew t >>= (fun r => k r t). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE bindA. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= [in RHS]bindE /=. +rewrite nth_error_rcons_size. +rewrite /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. +Qed. + +Let cnewchk T (s : coq_type T) (A : UU0) (k : coq_type T -> loc T -> M A) : + cnew s >>= (fun r => cchk r >> k s r) = cnew s >>= k s. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite nth_error_rcons_size. +rewrite /coerce. +by case: ml_type_eq_dec. +Qed. + +Let cchknewC T1 T2 (r : loc T1) (s : coq_type T2) (A : UU0) + (k : coq_type T2 -> loc T2 -> M A) : + cchk r >> (cnew s >>= fun r' => cchk r >> k s r') = cchk r >> (cnew s >>= k s). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +case Hc: (coerce T1 u) => [u'|] //. +rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /= bindA. +rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /= bindA. +rewrite (nth_error_rcons_some _ Hr) Hc. +do! rewrite bindE /=. +by rewrite /bindS /= MS_mapE /= fmapE bindA. +Qed. + +Let cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type T2 -> M A) : + cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. + rewrite bindfailf. + case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; + last by rewrite bindfailf. + case Hc: (coerce T2 v2) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1 bindfailf. +case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. + rewrite bindfailf. + case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr2 bindfailf. +case Hc: (coerce T1 v1) => [u1|]; last first. + rewrite bindfailf. + case Hc2: (coerce T2 v2) => [u|]; last by rewrite bindfailf. + rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1 Hc. + by rewrite bindfailf. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite Hr2. +case Hc2: (coerce T2 v2) => [u2|]; last by rewrite !bindfailf. +do! rewrite bindE /=. +by rewrite /bindS /= MS_mapE /= fmapE Hr1 Hc. +Qed. + +Let cchkget T (r : loc T) (A: UU0) (k : coq_type T -> M A) : + cchk r >> (cget r >>= k) = cget r >>= k. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +case Hc: (coerce T u) => [u'|] //. +do! rewrite bindE /=. +by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. +Qed. + +Let cgetchk T (r : loc T) (A: UU0) (k : coq_type T -> M A) : + cget r >>= (fun s => cchk r >> k s) = cget r >>= k. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +case Hc: (coerce T u) => [u'|] //. +do! rewrite bindE /=. +by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. +Qed. + +Let cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : + cchk r1 >> cput r2 s = cput r2 s >> cchk r1. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. + rewrite bindfailf. + case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; + last by rewrite bindfailf. + case Hc: (coerce T2' s) => [u|]; last by rewrite bindfailf. + rewrite bindE /= /bindS /= bindE /=. + by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). +case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. + rewrite bindfailf. + case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= Hr2. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec T1' T1) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec T2 T2') => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE/= bindE/=. + case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. + rewrite -Hr nth_error_set_nth /coerce. + move: Hr2; rewrite -Hr Hr1 => -[] HT2. + subst T1'. + by case: ml_type_eq_dec. + rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. + by case: ml_type_eq_dec. +rewrite bindE /= bindE /= Hr2 {1}/coerce. +case: ml_type_eq_dec => HT2 //. +subst T1' T2'. +rewrite -!eq_rect_eq. +rewrite bindE /= bindE /=. +case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. + rewrite -Hr nth_error_set_nth /coerce. + move: Hr2; rewrite -Hr Hr1 => -[] HT2. + subst T2. + by case: ml_type_eq_dec. +rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. +by case: ml_type_eq_dec. +Qed. + +Let cchkput T (r : loc T) (s : coq_type T) : + cchk r >> cput r s = cput r s. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +rewrite /coerce. +case: ml_type_eq_dec => HT; last by case: ml_type_eq_dec => // /esym //. +subst T'. +case: ml_type_eq_dec => // HT. +rewrite -eq_rect_eq {HT}. +do! rewrite bindE /=. +rewrite Hr. +case: ml_type_eq_dec => // HT. +by rewrite -eq_rect_eq. +Qed. + +Let cputchk T (r : loc T) (s : coq_type T) : + cput r s >> cchk r = cput r s. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +rewrite /coerce. +case: ml_type_eq_dec => HT //. +subst T'. +rewrite -!eq_rect_eq. +do! rewrite bindE /=. +rewrite nth_error_set_nth. +by case: ml_type_eq_dec. +Qed. + Let cputput T (r : loc T) (s s' : coq_type T) : - bind (cput r s) (fun _ => cput r s') = cput r s'. + cput r s >> cput r s' = cput r s'. Proof. congr mkActo. apply/boolp.funext => st. case: r s s' => {}T n s s' /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hst : (List.nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. rewrite {1 3}/coerce. case: ml_type_eq_dec => H /=; last by rewrite bindfailf. subst T'. rewrite !bindretf /= nth_error_set_nth /coerce. case: ml_type_eq_dec => // H. rewrite -eq_rect_eq. -set b := mkbind _ _. -set b' := mkbind _ _. -admit. -Admitted. +by rewrite set_set_nth eqxx. +Qed. Let cputget T (r : loc T) (s : coq_type T) (A : UU0) (k : coq_type T -> M A) : - bind (cput r s) (fun=> bind (cget r) k) = bind (cput r s) (fun=> k s). + cput r s >> (cget r >>= k) = cput r s >> k s. Proof. congr mkActo. apply/boolp.funext => st /=. case: r s k => {}T n s k /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hst : (List.nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. rewrite {1 3}/coerce. case: ml_type_eq_dec => H /=; last by rewrite bindfailf. subst T'. @@ -1773,28 +2033,189 @@ case: ml_type_eq_dec => H /=; last by rewrite bindfailf. by rewrite -eq_rect_eq !bindretf. Qed. -(* - cgetputchk : forall T (r : loc T), cget r >>= cput r = cchk r ; - cgetget : - forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), - cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s ; - cgetC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) - (k : coq_type M T1 -> coq_type M T2 -> M A), - cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = - cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)) ; - cputC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) - (s2 : coq_type M T2) (A : UU0), - loc_id r1 != loc_id r2 \/ mkbind T1 s1 = mkbind T2 s2 -> - cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1 ; - cputgetC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) - (A : UU0) (k : coq_type M T2 -> M A), - loc_id r1 != loc_id r2 -> - cput r1 s1 >> cget r2 >>= k = - cget r2 >>= (fun v => cput r1 s1 >> k v) }. -*) +Let cgetputchk T (r : loc T) : cget r >>= cput r = cchk r. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite /coerce. +case: (ml_type_eq_dec T1 T) => H //. +subst T1. +rewrite bindE /= bindE /= Hr. +case: (ml_type_eq_dec T T) => H //. +by rewrite -eq_rect_eq nth_error_set_nth_id. +Qed. + +Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : + cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite /coerce. +case: (ml_type_eq_dec T1 T) => H //. +subst T1. +rewrite bindE /= bindE /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite Hr. +case: (ml_type_eq_dec T T) => H //. +by rewrite -eq_rect_eq. +Qed. + +Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) + (k : coq_type T1 -> coq_type T2 -> M A) : + cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = + cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1: (nth_error _ _) => [[T1' u]|]; last first. + rewrite bindfailf. + case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. + case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. + rewrite !bindE /=. + by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. +case Hr2: (nth_error _ _) => [[T2' v]|]; last first. + rewrite bindfailf. + case Hc: (coerce T1 u) => [v|]; last by rewrite bindfailf. + rewrite !bindE /=. + by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr2 bindfailf. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec _ _) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE /=. + rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr1 {1}/coerce. + case: (ml_type_eq_dec _ _) => //; by rewrite bindfailf. +subst T1'. +rewrite -eq_rect_eq. +rewrite bindE /=. +rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr2 {1}/coerce. +case: (ml_type_eq_dec T2' T2) => HT2; last by rewrite !bindfailf. +subst T2'. +rewrite -eq_rect_eq. +rewrite !bindE /=. +rewrite !bindE /=. +rewrite /bindS MS_mapE /= fmapE /= bindA /= Hr1 /coerce. +case: (ml_type_eq_dec T1 T1) => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) + (s2 : coq_type T2) (A : UU0) : + loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> + cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. +Proof. +move=> H. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1: (nth_error _ _) => [[T1' u]|]; last first. + rewrite bindfailf. + case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. + case Hc: (coerce T2' s2) => [u|]; last by rewrite bindfailf. + rewrite !bindE /=. + rewrite !bindE /=. + case/boolP: (loc_id r1 == loc_id r2) => Hr. + by rewrite -(eqP Hr) Hr1 in Hr2. + by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). +case/boolP: (loc_id r1 == loc_id r2) H => [/eqP|] Hr /= H. + rewrite -Hr Hr1. + case/boolP: (T1 == T2) => /eqP HT; last first. + rewrite /coerce. + case: (ml_type_eq_dec T1 T1') => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec T2 T1') => HT2; last by rewrite bindfailf. + subst T1'. + rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. + by case: ml_type_eq_dec. + subst T1'. + rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. + by case: ml_type_eq_dec => [/esym|] //. + subst T2. + rewrite /coerce. + case: ml_type_eq_dec => HT1 //. + subst T1'. + rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /= !nth_error_set_nth. + case: ml_type_eq_dec => HT1 //. + rewrite -!eq_rect_eq. + case: H => // H. + by move: (JMeq_eq H) => <-. +case Hr2: (nth_error _ _) => [[T2' v]|]; last first. + rewrite bindfailf. + case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. + rewrite !bindE /=. + rewrite !bindE /=. + by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1). +rewrite {1 3}/coerce. +case: (ml_type_eq_dec _ _) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE /=. + rewrite !bindE /=. + rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. + by case: ml_type_eq_dec. +subst T1'. +rewrite -eq_rect_eq /coerce. +rewrite bindE /= bindE /= (nth_error_set_nth_other _ _ _ Hr2); + last by rewrite eq_sym. +case: ml_type_eq_dec => HT2' //. +subst T2'. +rewrite -eq_rect_eq bindE /= bindE /= (nth_error_set_nth_other _ _ Hr Hr1). +case: ml_type_eq_dec => // HT1. +by rewrite -eq_rect_eq set_set_nth (negbTE Hr). +Qed. + +Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) + (A : UU0) (k : coq_type T2 -> M A) : + loc_id r1 != loc_id r2 -> + cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). +Proof. +move=> Hr. +congr mkActo. +apply/boolp.funext => st /=. +rewrite !bindA. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1: (nth_error _ _) => [[T1' u]|]; last first. + rewrite bindfailf. + case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. + case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. + by rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. +case Hr2: (nth_error _ _) => [[T2' v]|]; last first. + rewrite bindfailf. + case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. + rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE. + by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1) bindfailf. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec _ _) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE Hr1 /coerce. + by case: ml_type_eq_dec. +subst T1'. +rewrite -eq_rect_eq /coerce. +rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite (nth_error_set_nth_other _ _ _ Hr2); last by rewrite eq_sym. +case: ml_type_eq_dec => HT2' //. +subst T2'. +rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /=. +rewrite /bindS /= MS_mapE /= fmapE Hr1. +case: ml_type_eq_dec => // HT1. +by rewrite -eq_rect_eq bindE. +Qed. End ModelTypedStore. (* TODO? From 9db04d14a84cecf8b10b1d021d7297fd310782c9 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Sat, 28 Jan 2023 18:36:43 +0900 Subject: [PATCH 05/82] more laws --- hierarchy.v | 4 ++++ monad_model.v | 37 +++++++++++++++++++++++++++++++++++++ 2 files changed, 41 insertions(+) diff --git a/hierarchy.v b/hierarchy.v index c21682d5..af02ed57 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1154,6 +1154,10 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cchk r >> cput r s = cput r s ; cputchk : forall T (r : loc T) (s : coq_type M T), cput r s >> cchk r = cput r s ; + cchkC : forall T1 T2 (r1: loc T1) (r2: loc T2), + cchk r1 >> cchk r2 = cchk r2 >> cchk r1 ; + cchkdup : forall T (r : loc T), + cchk r >> cchk r = cchk r ; }. #[short(type=typedStoreMonad)] diff --git a/monad_model.v b/monad_model.v index b0f841aa..5384a3d1 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1861,6 +1861,43 @@ do! rewrite bindE /=. by rewrite /bindS /= MS_mapE /= fmapE bindA. Qed. +Let cchkC T1 T2 (r1: loc T1) (r2: loc T2) : + cchk r1 >> cchk r2 = cchk r2 >> cchk r1. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. + rewrite bindfailf. + case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; + last by rewrite bindfailf. + case Hc: (coerce T2 v2) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= Hr1. +case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. + rewrite bindfailf. + case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= Hr2. +case Hc: (coerce T1 v1) => [u1|]; last first. + rewrite bindfailf. + case Hc2: (coerce T2 v2) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= Hr1 Hc. +rewrite bindE /= bindE /= Hr2. +case Hc2: (coerce T2 v2) => [u2|]; last by rewrite !bindfailf. +do! rewrite bindE /=. +by rewrite Hr1 Hc. +Qed. + +Let cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1 : (nth_error st (loc_id r)) => [[T1' v1]|] //=. +case Hc: (coerce T v1) => [u1|] //. +by rewrite bindE /= bindE /= Hr1 Hc. +Qed. + Let cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type T2 -> M A) : cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). Proof. From edde048419e739817282807ea3f9cdc9dd0703ee Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Mon, 30 Jan 2023 12:35:15 +0900 Subject: [PATCH 06/82] add crun and example_typed_store.v (fact_ref) --- _CoqProject | 1 + example_typed_store.v | 93 +++++++++++++++++++++++++++++++++++++++++++ hierarchy.v | 9 ++++- monad_model.v | 17 ++++++++ 4 files changed, 118 insertions(+), 2 deletions(-) create mode 100644 example_typed_store.v diff --git a/_CoqProject b/_CoqProject index 66bb9250..91413f8b 100644 --- a/_CoqProject +++ b/_CoqProject @@ -26,6 +26,7 @@ example_relabeling.v example_quicksort.v example_iquicksort.v example_monty.v +example_typed_store.v smallstep.v example_transformer.v category_ext.v diff --git a/example_typed_store.v b/example_typed_store.v new file mode 100644 index 00000000..4123112d --- /dev/null +++ b/example_typed_store.v @@ -0,0 +1,93 @@ +(* monae: Monadic equational reasoning in Coq *) +(* Copyright (C) 2020 monae authors, license: LGPL-2.1-or-later *) +Require Import ZArith. +From mathcomp Require Import all_ssreflect. +From mathcomp Require boolp. +From infotheo Require Import ssrZ. +Require Import monae_lib hierarchy. (*monad_lib fail_lib state_lib.*) + +(******************************************************************************) +(* Type store examples *) +(* *) +(******************************************************************************) + +Local Open Scope monae_scope. + +Inductive ml_type : Set := + | ml_int + | ml_bool + | ml_unit + | ml_ref (_ : ml_type) + | ml_arrow (_ : ml_type) (_ : ml_type) + | ml_undef. + +Inductive undef_t : Set := Undef. + +Module MLtypes. +Definition ml_type_eq_dec (T1 T2 : ml_type) : {T1=T2}+{T1<>T2}. +revert T2; induction T1; destruct T2; + try (right; intro; discriminate); try (now left); + try (case (IHT1_5 T2_5); [|right; injection; intros; contradiction]); + try (case (IHT1_4 T2_4); [|right; injection; intros; contradiction]); + try (case (IHT1_3 T2_3); [|right; injection; intros; contradiction]); + try (case (IHT1_2 T2_2); [|right; injection; intros; contradiction]); + (case (IHT1 T2) || case (IHT1_1 T2_1)); try (left; now subst); + right; injection; intros; contradiction. +Defined. + +Local Definition ml_type := ml_type. +Local Definition undef := Undef. +Variant loc : ml_type -> Type := mkloc T : nat -> loc T. + +Section with_monad. +Context [M : Type -> Type]. + +Local (* Generated type translation function *) +Fixpoint coq_type (T : ml_type) : Type := + match T with + | ml_int => nat + | ml_bool => bool + | ml_unit => unit + | ml_arrow T1 T2 => coq_type T1 -> M (coq_type T2) + | ml_ref T1 => loc T1 + | ml_undef => undef_t + end. +End with_monad. +Local Definition ml_undef := ml_undef. +End MLtypes. + +Module IMonadTS := MonadTypedStore (MLtypes). +Import MLtypes. +Import IMonadTS. + +Lemma cgetret (M : typedStoreMonad) T : + cget T = fun r => (cget _ r : M _) >>= Ret. +Proof. by apply boolp.funext => r; rewrite bindmret. Qed. + +Section factorial. +Variable M : typedStoreMonad. +Local Definition coq_type := @MLtypes.coq_type M. + +Fixpoint fact_ref (r : loc ml_int) (n : nat) : M unit := + if n is m.+1 then cget _ r >>= fun p => cput _ r (n * p) >> fact_ref r m + else Ret tt. + +Theorem fact_ref_ok n : + crun _ (cnew ml_int 1 >>= fun r => fact_ref r n >> cget _ r) = + Some (fact_rec n). +Proof. +set fn := fact_rec n. +set m := n. +set s := 1. +have smn : s * fact_rec m = fn by rewrite mul1n. +elim: m s smn => [|m IH] s /= smn. + rewrite /fact_ref -smn muln1. + under [fun r => _]boolp.funext do rewrite bindretf. + by rewrite cgetret cnewget crunret // crunnew. +under [fun r => _]boolp.funext do rewrite bindA. +rewrite cnewget. +under [fun r => _]boolp.funext do rewrite bindA. +rewrite (cnewput ml_int s _ _ (fun r s => fact_ref r m >> cget ml_int r)). +by rewrite IH // (mulnC m.+1) -mulnA. +Qed. +End factorial. diff --git a/hierarchy.v b/hierarchy.v index af02ed57..6e0035be 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1099,8 +1099,9 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cget : forall {T}, loc T -> M (coq_type M T) ; cput : forall {T}, loc T -> coq_type M T -> M unit ; cchk : forall {T}, loc T -> M unit ; - cnewget : forall T (s : coq_type M T) A (k : coq_type M T -> M A), - cnew s >>= (fun r => cget r >>= k) = cnew s >> k s ; + crun : forall {A : UU0}, M A -> option A ; (* execute in empty store *) + cnewget : forall T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A), + cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s) ; cnewgetC : forall T T' (r : loc T) (s : coq_type M T') A (k : loc T' -> coq_type M T -> M A), @@ -1158,6 +1159,10 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cchk r1 >> cchk r2 = cchk r2 >> cchk r1 ; cchkdup : forall T (r : loc T), cchk r >> cchk r = cchk r ; + crunret : forall (A B : UU0) (m : M A) (s : B), + crun m -> crun (m >> Ret s) = Some s ; + crunnew : forall T (s : coq_type M T), + crun (cnew s) ; }. #[short(type=typedStoreMonad)] diff --git a/monad_model.v b/monad_model.v index 5384a3d1..8a8d6870 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1713,6 +1713,12 @@ Definition cchk T (r : loc T) : M unit := if coerce T u is Some _ then Ret (tt, st) else fail else fail). +Definition crun (A : UU0) (m : M A) : option A := + match ofActo m nil with + | inl _ => None + | inr (a, _) => Some a + end. + Let ret : forall A, idfun A -> M A := fun A a => mkActo (Ret a). Let bind A B (m : M A) (f : A -> M B) : M B := mkActo (ofActo m >>= (fun a => ofActo (f a))). @@ -2253,6 +2259,17 @@ rewrite /bindS /= MS_mapE /= fmapE Hr1. case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq bindE. Qed. + +Let crunret (A B : UU0) (m : M A) (s : B) : + crun m -> crun (m >> ret s) = Some s. +Proof. +case: m => T /= m. +rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +by case Hm: (m [::]). +Qed. + +Let crunnew T (s : coq_type T) : crun (cnew s). +Proof. by []. Qed. End ModelTypedStore. (* TODO? From 1a23435e329d8b8b1d2750e81ddd55364a5b5e00 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Mon, 30 Jan 2023 14:48:50 +0900 Subject: [PATCH 07/82] add fibonacci example with two refs --- example_typed_store.v | 92 +++++++++++++++++++++++++++++--- hierarchy.v | 11 +++- impredicative_set/imonad_model.v | 2 + monad_model.v | 6 +++ 4 files changed, 102 insertions(+), 9 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 4123112d..1a02f61e 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -60,21 +60,28 @@ Module IMonadTS := MonadTypedStore (MLtypes). Import MLtypes. Import IMonadTS. +Arguments cget {s} [T]. +Arguments cput {s} [T]. +Arguments cchk {s} [T]. +Arguments crun {s} [A]. + Lemma cgetret (M : typedStoreMonad) T : - cget T = fun r => (cget _ r : M _) >>= Ret. + @cget _ T = fun r => (cget r : M _) >>= Ret. Proof. by apply boolp.funext => r; rewrite bindmret. Qed. +Lemma crunnew0 (M : typedStoreMonad) T s : crun (cnew T s : M _). +Proof. by rewrite -[cnew _ _]bindskipf crunnew // crunskip. Qed. + Section factorial. Variable M : typedStoreMonad. -Local Definition coq_type := @MLtypes.coq_type M. +Notation coq_type := (@MLtypes.coq_type M). Fixpoint fact_ref (r : loc ml_int) (n : nat) : M unit := - if n is m.+1 then cget _ r >>= fun p => cput _ r (n * p) >> fact_ref r m + if n is m.+1 then cget r >>= fun p => cput r (n * p) >> fact_ref r m else Ret tt. Theorem fact_ref_ok n : - crun _ (cnew ml_int 1 >>= fun r => fact_ref r n >> cget _ r) = - Some (fact_rec n). + crun (cnew ml_int 1 >>= fun r => fact_ref r n >> cget r) = Some (fact_rec n). Proof. set fn := fact_rec n. set m := n. @@ -83,11 +90,82 @@ have smn : s * fact_rec m = fn by rewrite mul1n. elim: m s smn => [|m IH] s /= smn. rewrite /fact_ref -smn muln1. under [fun r => _]boolp.funext do rewrite bindretf. - by rewrite cgetret cnewget crunret // crunnew. + by rewrite cgetret cnewget crunret // crunnew0. under [fun r => _]boolp.funext do rewrite bindA. rewrite cnewget. under [fun r => _]boolp.funext do rewrite bindA. -rewrite (cnewput ml_int s _ _ (fun r s => fact_ref r m >> cget ml_int r)). +rewrite (cnewput ml_int s _ _ (fun r s => fact_ref r m >> cget r)). by rewrite IH // (mulnC m.+1) -mulnA. Qed. End factorial. + +Section fibonacci. +Variable M : typedStoreMonad. +Notation coq_type := (@MLtypes.coq_type M). + +Fixpoint fibo_rec n := + if n is m.+1 then + if m is k.+1 then fibo_rec k + fibo_rec m else 1 + else 1. + +Fixpoint fibo_ref n (a b : loc ml_int) : M unit := + if n is n.+1 then + cget a >>= (fun x => cget b >>= fun y => cput a y >> cput b (x + y)) + >> fibo_ref n a b + else skip. + +Theorem fibo_ref_ok n : + crun (cnew ml_int 1 >>= + (fun a => cnew ml_int 1 >>= fun b => fibo_ref n a b >> cget a)) + = Some (fibo_rec n). +Proof. +set x := 1. +pose y := x. +rewrite -{2}/y. +pose i := n. +rewrite -{1}/i. +have : (x, y) = (fibo_rec (n-i), fibo_rec (n.+1-i)). + by rewrite subnn -addn1 addKn. +have : i <= n by []. +elim: i x y => [|i IH] x y Hi. + rewrite !subn0 => -[] -> ->. + rewrite -/(fibo_rec n.+1). + under [fun a => _]boolp.funext do + under [fun b => _]boolp.funext do rewrite [fibo_ref _ _ _]/= bindskipf. + under [fun a => _]boolp.funext do rewrite cgetret. + set f := fun a => _. + rewrite -(cnewchk ml_int (fibo_rec n) _ (fun => f)) {}/f. + under [fun b => _]boolp.funext do rewrite cnewgetC. + rewrite cnewget. + by rewrite -bindA crunret // crunnew // crunnew0. +rewrite subSS => -[] Hx Hy. +rewrite -(IH y (x + y) (ltnW Hi)); last first. + subst x y. + congr pair. + case: n Hi {IH} => // n. + rewrite ltnS => Hi. + rewrite subSS. + by rewrite -addn2 -addn1 -addnBAC // -addnBAC // addn2 addn1. +rewrite /=. +under [fun a => _]boolp.funext do +under [fun a => _]boolp.funext do rewrite !bindA. +rewrite -cnewchk. +under [fun b => _]boolp.funext do rewrite cnewgetC. +rewrite cnewget. +under [fun a => _]boolp.funext do +under [fun a => _]boolp.funext do rewrite !bindA. +under [fun a => _]boolp.funext do rewrite cnewget. +under [fun a => _]boolp.funext do +under [fun a => _]boolp.funext do rewrite !bindA. +set f := fun a => _. +rewrite -(cnewchk ml_int _ _ (fun => f)) {}/f. +under [fun a => _]boolp.funext do rewrite cnewputC. +set f := fun r (_ : nat) => + (cnew ml_int y >>= + (fun r' : loc ml_int => + cput r' (x + y) >> (fibo_ref i r r' >> cget r))). +rewrite (cnewput _ _ _ _ f) {}/f. +set f := fun r r' (_ : nat) => fibo_ref i r r' >> cget r. +by under [fun a => _]boolp.funext do rewrite (cnewput _ _ _ _ (f _)) {}/f. +Qed. +End fibonacci. diff --git a/hierarchy.v b/hierarchy.v index 6e0035be..e923c5c7 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1109,6 +1109,11 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cget r >>= (fun u => cnew s >>= (fun r' => k r' u)) ; cnewput : forall T (s t : coq_type M T) A (k : loc T -> coq_type M T -> M A), cnew s >>= (fun r => cput r t >> k r t) = cnew t >>= (fun r => k r t) ; + cnewputC : + forall T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A + (k : loc T' -> M A), + cchk r >> (cnew s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew s' >>= k) ; cputput : forall T (r : loc T) (s s' : coq_type M T), cput r s >> cput r s' = cput r s' ; cputget : @@ -1161,8 +1166,10 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cchk r >> cchk r = cchk r ; crunret : forall (A B : UU0) (m : M A) (s : B), crun m -> crun (m >> Ret s) = Some s ; - crunnew : forall T (s : coq_type M T), - crun (cnew s) ; + crunskip : + crun skip = Some tt ; + crunnew : forall (A : UU0) T (m : M A) (s : coq_type M T), + crun m -> crun (m >> cnew s) ; }. #[short(type=typedStoreMonad)] diff --git a/impredicative_set/imonad_model.v b/impredicative_set/imonad_model.v index a422caae..1d8198f2 100644 --- a/impredicative_set/imonad_model.v +++ b/impredicative_set/imonad_model.v @@ -1173,8 +1173,10 @@ Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. Let associative : BindLaws.associative bind. Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. +(* HB.instance Definition xyz := isMonad_ret_bind.Build M left_neutral right_neutral associative. +*) Lemma nth_error_set_nth T (x : T) st n : List.nth_error (set_nth x st n x) n = Some x. diff --git a/monad_model.v b/monad_model.v index 8a8d6870..a413d568 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1837,6 +1837,12 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. Qed. +Let cnewputC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A + (k : loc T' -> M A) : + cchk r >> (cnew s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew s' >>= k). +Admitted. + Let cnewchk T (s : coq_type T) (A : UU0) (k : coq_type T -> loc T -> M A) : cnew s >>= (fun r => cchk r >> k s r) = cnew s >>= k s. Proof. From 183f368cbfb6d562d026a62ec473757e27a38649 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Mon, 30 Jan 2023 15:20:16 +0900 Subject: [PATCH 08/82] adjust laws and finish proofs --- example_typed_store.v | 21 ++++++--------------- hierarchy.v | 15 +++++++-------- monad_model.v | 42 +++++++++++++++++++++++++++++++++++------- 3 files changed, 48 insertions(+), 30 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 1a02f61e..706a7c01 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -94,8 +94,7 @@ elim: m s smn => [|m IH] s /= smn. under [fun r => _]boolp.funext do rewrite bindA. rewrite cnewget. under [fun r => _]boolp.funext do rewrite bindA. -rewrite (cnewput ml_int s _ _ (fun r s => fact_ref r m >> cget r)). -by rewrite IH // (mulnC m.+1) -mulnA. +by rewrite cnewput IH // (mulnC m.+1) -mulnA. Qed. End factorial. @@ -133,11 +132,9 @@ elim: i x y => [|i IH] x y Hi. under [fun a => _]boolp.funext do under [fun b => _]boolp.funext do rewrite [fibo_ref _ _ _]/= bindskipf. under [fun a => _]boolp.funext do rewrite cgetret. - set f := fun a => _. - rewrite -(cnewchk ml_int (fibo_rec n) _ (fun => f)) {}/f. + rewrite -cnewchk. under [fun b => _]boolp.funext do rewrite cnewgetC. - rewrite cnewget. - by rewrite -bindA crunret // crunnew // crunnew0. + by rewrite cnewget -bindA crunret // crunnew // crunnew0. rewrite subSS => -[] Hx Hy. rewrite -(IH y (x + y) (ltnW Hi)); last first. subst x y. @@ -157,15 +154,9 @@ under [fun a => _]boolp.funext do rewrite !bindA. under [fun a => _]boolp.funext do rewrite cnewget. under [fun a => _]boolp.funext do under [fun a => _]boolp.funext do rewrite !bindA. -set f := fun a => _. -rewrite -(cnewchk ml_int _ _ (fun => f)) {}/f. +rewrite -cnewchk. under [fun a => _]boolp.funext do rewrite cnewputC. -set f := fun r (_ : nat) => - (cnew ml_int y >>= - (fun r' : loc ml_int => - cput r' (x + y) >> (fibo_ref i r r' >> cget r))). -rewrite (cnewput _ _ _ _ f) {}/f. -set f := fun r r' (_ : nat) => fibo_ref i r r' >> cget r. -by under [fun a => _]boolp.funext do rewrite (cnewput _ _ _ _ (f _)) {}/f. +rewrite cnewput. +by under [fun a => _]boolp.funext do rewrite cnewput. Qed. End fibonacci. diff --git a/hierarchy.v b/hierarchy.v index e923c5c7..1abd4c92 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1107,8 +1107,8 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) (k : loc T' -> coq_type M T -> M A), cchk r >> (cnew s >>= (fun r' => cget r >>= k r')) = cget r >>= (fun u => cnew s >>= (fun r' => k r' u)) ; - cnewput : forall T (s t : coq_type M T) A (k : loc T -> coq_type M T -> M A), - cnew s >>= (fun r => cput r t >> k r t) = cnew t >>= (fun r => k r t) ; + cnewput : forall T (s t : coq_type M T) A (k : loc T -> M A), + cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k ; cnewputC : forall T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A (k : loc T' -> M A), @@ -1140,13 +1140,12 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v) ; cnewchk : - forall T (s : coq_type M T) (A : UU0) (k : coq_type M T -> loc T -> M A), - cnew s >>= (fun r => cchk r >> k s r) = cnew s >>= k s ; + forall T (s : coq_type M T) (A : UU0) (k : loc T -> M A), + cnew s >>= (fun r => cchk r >> k r) = cnew s >>= k ; cchknewC : - forall T1 T2 (r : loc T1) (s : coq_type M T2) (A : UU0) - (k : coq_type M T2 -> loc T2 -> M A), - cchk r >> (cnew s >>= fun r' => cchk r >> k s r') = - cchk r >> (cnew s >>= k s) ; + forall T1 T2 (r : loc T1) (s : coq_type M T2) (A : UU0) (k : loc T2 -> M A), + cchk r >> (cnew s >>= fun r' => cchk r >> k r') = + cchk r >> (cnew s >>= k) ; cchkgetC : forall T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type M T2 -> M A), cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s) ; diff --git a/monad_model.v b/monad_model.v index a413d568..bffd38e7 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1772,10 +1772,19 @@ Lemma nth_error_set_nth_none A (def : A) st m n a b : nth_error (set_nth def st n b) m = None. Proof. by elim: m st n => [|m IH] [|c st] [|n] //=; apply IH. Qed. +Lemma nth_error_size_set_nth T def (st : seq T) n a b : + nth_error st n = Some a -> size (set_nth def st n b) = size st. +Proof. by elim: n st => [|n IH] [|c st] //= /IH ->. Qed. + Lemma set_nth_rcons A (def : A) st a b : set_nth def (rcons st a) (size st) b = rcons st b. Proof. by elim: st => //= c st ->. Qed. +Lemma nth_error_set_nth_rcons T def (st : seq T) n a b c : + nth_error st n = Some a -> + set_nth def (rcons st c) n b = rcons (set_nth def st n b) c. +Proof. by elim: n st => [|n IH] [|d st] //= /IH ->. Qed. + (* Make ml_type an eqType *) Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. Lemma ml_type_eqP : Equality.axiom ml_type_eqb. @@ -1822,8 +1831,8 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. Qed. -Let cnewput T (s t : coq_type T) A (k : loc T -> coq_type T -> M A) : - cnew s >>= (fun r => cput r t >> k r t) = cnew t >>= (fun r => k r t). +Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : + cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. congr mkActo. apply/boolp.funext => st /=. @@ -1841,10 +1850,29 @@ Let cnewputC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A (k : loc T' -> M A) : cchk r >> (cnew s' >>= fun r' => cput r s >> k r') = cput r s >> (cnew s' >>= k). -Admitted. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. +subst T1. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite (nth_error_rcons_some _ Hr). +rewrite /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /=. +rewrite (nth_error_size_set_nth _ _ Hr). +by rewrite (nth_error_set_nth_rcons _ _ _ Hr). +Qed. -Let cnewchk T (s : coq_type T) (A : UU0) (k : coq_type T -> loc T -> M A) : - cnew s >>= (fun r => cchk r >> k s r) = cnew s >>= k s. +Let cnewchk T (s : coq_type T) (A : UU0) (k : loc T -> M A) : + cnew s >>= (fun r => cchk r >> k r) = cnew s >>= k. Proof. congr mkActo. apply/boolp.funext => st /=. @@ -1857,8 +1885,8 @@ by case: ml_type_eq_dec. Qed. Let cchknewC T1 T2 (r : loc T1) (s : coq_type T2) (A : UU0) - (k : coq_type T2 -> loc T2 -> M A) : - cchk r >> (cnew s >>= fun r' => cchk r >> k s r') = cchk r >> (cnew s >>= k s). + (k : loc T2 -> M A) : + cchk r >> (cnew s >>= fun r' => cchk r >> k r') = cchk r >> (cnew s >>= k). Proof. congr mkActo. apply/boolp.funext => st /=. From c3725c8f30be49ec99d11e11bb74967a26682d9b Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Mon, 30 Jan 2023 15:40:37 +0900 Subject: [PATCH 09/82] test: cnew_get_def --- hierarchy.v | 10 ++++++++-- monad_model.v | 5 ++--- 2 files changed, 10 insertions(+), 5 deletions(-) diff --git a/hierarchy.v b/hierarchy.v index 1abd4c92..2b2720dd 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1093,6 +1093,13 @@ Arguments mkbind {M}. Definition loc_id {T} (r : loc T) : nat := let: mkloc _ n := r in n. +Definition cnewget_def (M : UU0 -> UU0) + (cnew : forall {T}, coq_type M T -> M (loc T)) + (cget : forall {T}, loc T -> M (coq_type M T)) + (bind : forall {A B}, M A -> (A -> M B) -> M B) + := forall T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A), + bind (cnew s) (fun r => bind (cget r) (k r)) = bind (cnew s) (fun r => k r s). + HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) of Monad M := { cnew : forall {T}, coq_type M T -> M (loc T) ; @@ -1100,8 +1107,7 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cput : forall {T}, loc T -> coq_type M T -> M unit ; cchk : forall {T}, loc T -> M unit ; crun : forall {A : UU0}, M A -> option A ; (* execute in empty store *) - cnewget : forall T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A), - cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s) ; + cnewget : @cnewget_def M (@cnew) (@cget) (@bind [the monad of M]) ; cnewgetC : forall T T' (r : loc T) (s : coq_type M T') A (k : loc T' -> coq_type M T -> M A), diff --git a/monad_model.v b/monad_model.v index bffd38e7..818d15e7 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1796,10 +1796,9 @@ Definition ml_type_eq_mixin := EqMixin ml_type_eqP. Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. (* Prove the laws *) -Let cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : - cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). +Let cnewget : cnewget_def cnew cget bind. Proof. -congr mkActo. +move=> T s A k; congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. From 71630573637629e3972ee7bace0bf004d1070b1f Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 31 Jan 2023 01:58:55 +0900 Subject: [PATCH 10/82] build model using Unset Universe Checking --- hierarchy.v | 12 +--- monad_model.v | 169 +++++++++++++++++++++++++++++++++++++++++++++++--- 2 files changed, 165 insertions(+), 16 deletions(-) diff --git a/hierarchy.v b/hierarchy.v index 2b2720dd..b28ec49a 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1093,13 +1093,6 @@ Arguments mkbind {M}. Definition loc_id {T} (r : loc T) : nat := let: mkloc _ n := r in n. -Definition cnewget_def (M : UU0 -> UU0) - (cnew : forall {T}, coq_type M T -> M (loc T)) - (cget : forall {T}, loc T -> M (coq_type M T)) - (bind : forall {A B}, M A -> (A -> M B) -> M B) - := forall T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A), - bind (cnew s) (fun r => bind (cget r) (k r)) = bind (cnew s) (fun r => k r s). - HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) of Monad M := { cnew : forall {T}, coq_type M T -> M (loc T) ; @@ -1107,7 +1100,8 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cput : forall {T}, loc T -> coq_type M T -> M unit ; cchk : forall {T}, loc T -> M unit ; crun : forall {A : UU0}, M A -> option A ; (* execute in empty store *) - cnewget : @cnewget_def M (@cnew) (@cget) (@bind [the monad of M]) ; + cnewget : forall T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A), + cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s) ; cnewgetC : forall T T' (r : loc T) (s : coq_type M T') A (k : loc T' -> coq_type M T -> M A), @@ -1124,7 +1118,7 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cput r s >> cput r s' = cput r s' ; cputget : forall T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A), - cput r s >> cget r >>= k = cput r s >> k s ; + cput r s >> (cget r >>= k) = cput r s >> k s ; cgetputchk : forall T (r : loc T), cget r >>= cput r = cchk r ; cgetget : forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), diff --git a/monad_model.v b/monad_model.v index 818d15e7..01b949d9 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1729,20 +1729,134 @@ Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. Let associative : BindLaws.associative bind. Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. +Search "isMonad" "Store". +Print MonadTypedStore.MonadTypedStore_isMonadTypedStore_mixin. +Print MonadTypedStore.axioms_. +Print isMonadTypedStore.axioms_. +Print isFunctor.axioms_. +Print isMonad.axioms_. + +Let actm (A B : UU0) (f : A -> B) (m : M A) : M B := + mkActo (actm _ _ f (ofActo m)). + +Definition isFunctorTS : isFunctor.axioms_ M. +exists actm. +- move=> A. apply boolp.funext => -[] {A} A m. + by rewrite /actm functor_id. +- move=> A B C g h. apply boolp.funext => m. + case: m h => {A} A m h /=. + by rewrite /actm /= functor_o. +Defined. + +Search "Functor". +Canonical isFunctorTS. + +About Functor. +Print Functor.axioms_. + +Unset Universe Checking. +Canonical Structure FunctorTS : Functor M := Functor.Class isFunctorTS. +Canonical Structure functorTS : functor := Functor.Pack isFunctorTS. +Set Universe Checking. + +Let ret_naturality : naturality [the functor of idfun] M ret. +Proof. by []. Qed. + +Canonical Structure naturalityTSret := isNatural.Build _ _ _ ret_naturality. + +Let ret' : [the functor of idfun] ~> [the functor of M] := + @Nattrans.Pack _ _ ret (Nattrans.Class naturalityTSret). + +Unset Universe Checking. +Let join' : [the functor of M \o M] ~~> [the functor of M] := + fun _ m => bind m idfun. +Set Universe Checking. + +Let actm_bind (a b c : UU0) (f : a -> b) m (g : c -> M a) : + (actm f) (bind m g) = bind m (actm f \o g). +Proof. +congr mkActo. +case: m g => {}c m g /=. +rewrite bindE fmapE 2!bindA /=. +congr (_ >>= _). +Qed. + +Lemma mkActoK A (m : MS _ _ A) : ofActo (mkActo m) = m. +Proof. done. Qed. + +Unset Universe Checking. +Let join'_naturality : naturality _ _ join'. +Proof. +move=> a b h. +unfold join'. +rewrite /=. apply: boolp.funext => mm /=. +unfold hierarchy.actm, isFunctor.actm. +rewrite /= (actm_bind h mm idfun) /= /bind bindA. +apply f_equal. +apply f_equal. +apply boolp.funext => x /=. +rewrite /retS /= fmapE. +case: x h mm => {}a x h mm. +rewrite mkActoK /hierarchy.actm /= /actm /= fmapE !bindretf mkActoK. +reflexivity. +Qed. + +Canonical Structure naturalityTS := isNatural.Build _ _ _ join'_naturality. + +Let join : [the functor of M \o M] ~> [the functor of M] := + @Nattrans.Pack _ _ join' (Nattrans.Class naturalityTS). + + +Canonical Structure isMonadTS : isMonad.axioms_ M isFunctorTS. +exists ret' join bind. +- move=> A B f m. + rewrite /join /= /join' /bind. + apply f_equal. + rewrite bindA. + reflexivity. +- move=> A. + apply boolp.funext => x /=. + case: x => {}A m. + reflexivity. +- move=> A. + apply boolp.funext => x /=. + case: x => {}A m. + rewrite /join' /bind /=. + apply f_equal. + rewrite fmapE /ret. + rewrite bindA /=. + under [fun _ => _]boolp.funext do rewrite bindretf mkActoK. + rewrite bindmret. + reflexivity. +- move=> A /=. + rewrite /join' /bind. + apply boolp.funext => x /=. + apply f_equal. + rewrite fmapE. + rewrite !bindA. + apply f_equal. + apply boolp.funext => y /=. + rewrite bindretf mkActoK. + reflexivity. +Defined. + +Canonical Structure MonadTS : Monad M := Monad.Class isMonadTS. +Canonical Structure monadTS : monad := Monad.Pack isMonadTS. +Set Universe Checking. + (* Doesn't work HB.instance Definition xyz := isMonad_ret_bind.Build M left_neutral right_neutral associative. *) (* Since we couldn't build the instance, redefine notations *) -Local Notation "m >>= f" := (bind m f). -Local Notation "m >> f" := (bind m (fun=> f)). +(*Local Notation "m >>= f" := (bind m f). +Local Notation "m >> f" := (bind m (fun=> f)).*) (* Basic lemmas for standard library's nth_error *) Lemma nth_error_set_nth T (def x : T) st n : nth_error (set_nth def st n x) n = Some x. Proof. -elim: n st => [|z IH] [] //. -clear IH. +elim: n st => [|z IH] [] // {IH}. elim: z.+1 => [|n <-] //=. by rewrite set_nth_nil. Qed. @@ -1796,9 +1910,10 @@ Definition ml_type_eq_mixin := EqMixin ml_type_eqP. Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. (* Prove the laws *) -Let cnewget : cnewget_def cnew cget bind. +Let cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : + cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. -move=> T s A k; congr mkActo. +congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. @@ -2301,8 +2416,48 @@ rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. by case Hm: (m [::]). Qed. -Let crunnew T (s : coq_type T) : crun (cnew s). +Let crunskip : crun skip = Some tt. Proof. by []. Qed. + +Let crunnew (A : UU0) T (m : M A) (s : coq_type T) : + crun m -> crun (m >> cnew s). +Proof. +case: m => U /= m. +rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +by case Hm: (m [::]). +Qed. + +Unset Universe Checking. +Definition isMonadTypedStoreModel : + isMonadTypedStore.axioms_ M isFunctorTS isMonadTS. +Proof. +exists cnew cget cput cchk crun. +- exact cnewget. +- exact cnewgetC. +- exact cnewput. +- exact cnewputC. +- exact cputput. +- exact cputget. +- exact cgetputchk. +- exact cgetget. +- exact cgetC. +- exact cputC. +- exact cputgetC. +- exact cnewchk. +- exact cchknewC. +- exact cchkgetC. +- exact cchkget. +- exact cgetchk. +- exact cchkputC. +- exact cchkput. +- exact cputchk. +- exact cchkC. +- exact cchkdup. +- exact crunret. +- exact crunskip. +- exact crunnew. +Defined. +Set Universe Checking. End ModelTypedStore. (* TODO? From 67afd95995c8aa36a626818decbb77abd8c2f39a Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 31 Jan 2023 09:29:18 +0900 Subject: [PATCH 11/82] move recursive TypedStoreModel to typed_store_model.v and replace it by weaker version in monad_model.v --- _CoqProject | 1 + monad_model.v | 419 ++++++++---------------- typed_store_model.v | 762 ++++++++++++++++++++++++++++++++++++++++++++ 3 files changed, 901 insertions(+), 281 deletions(-) create mode 100644 typed_store_model.v diff --git a/_CoqProject b/_CoqProject index 91413f8b..752a4c65 100644 --- a/_CoqProject +++ b/_CoqProject @@ -20,6 +20,7 @@ category.v gcm_model.v altprob_model.v monad_transformer.v +typed_store_model.v example_spark.v example_nqueens.v example_relabeling.v diff --git a/monad_model.v b/monad_model.v index 01b949d9..32d4ced3 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1656,204 +1656,13 @@ End modelplusarray. End ModelPlusArray. HB.export ModelPlusArray. -Require Import JMeq. - -Module Type MLTYdef. -Include MLTY. -Parameter ml_undef : ml_type. -Parameter undef : forall M, coq_type M ml_undef. -End MLTYdef. +Section nth_error. +Context [T : Type] (def : T) (st : seq T). -Module ModelTypedStore (MLtypes : MLTYdef). -Module MTypedStore := MonadTypedStore (MLtypes). -Import MLtypes. -Import MTypedStore. - -#[bypass_check(positivity)] -Inductive acto : UU0 -> UU0 := - mkActo : forall T : UU0, - MS (seq (binding acto)) [the monad of option_monad] T -> acto T. -Local Notation M := acto. -Local Notation coq_type := (coq_type M). - -Definition ofActo T (m : M T) - : MS (seq (binding M)) [the monad of option_monad] T := - let: mkActo _ m' := m in m'. - -Definition def := mkbind ml_undef (undef M). +(* Basic lemmas for standard library's nth_error *) Local Notation nth_error := List.nth_error. -Definition cnew T (v : coq_type T) : M (loc T) := - mkActo (fun st => let n := size st in Ret (mkloc T n, rcons st (mkbind T v))). - -Definition coerce (T1 T2 : ml_type) (v : coq_type T1) : option (coq_type T2) := - match ml_type_eq_dec T1 T2 with - | left H => Some (eq_rect _ _ v _ H) - | right _ => None - end. - -Definition cget T (r : loc T) : M (coq_type T) := - mkActo (fun st => - if nth_error st (loc_id r) is Some (mkbind T' v) then - if coerce T v is Some u then Ret (u, st) else fail - else fail). - -Definition cput T (r : loc T) (v : coq_type T) : M unit := - mkActo (fun st => - let n := loc_id r in - if nth_error st n is Some (mkbind T' _) then - if coerce T' v is Some u then - Ret (tt, set_nth def st n (mkbind T' u)) - else fail - else fail). - -Definition cchk T (r : loc T) : M unit := - mkActo (fun st => - if nth_error st (loc_id r) is Some (mkbind T' u) then - if coerce T u is Some _ then Ret (tt, st) else fail - else fail). - -Definition crun (A : UU0) (m : M A) : option A := - match ofActo m nil with - | inl _ => None - | inr (a, _) => Some a - end. - -Let ret : forall A, idfun A -> M A := fun A a => mkActo (Ret a). -Let bind A B (m : M A) (f : A -> M B) : M B := - mkActo (ofActo m >>= (fun a => ofActo (f a))). -Let left_neutral : BindLaws.left_neutral bind ret. -Proof. by move=> A B a f; rewrite /bind /ret bindretf; case: (f a). Qed. -Let right_neutral : BindLaws.right_neutral bind ret. -Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. -Let associative : BindLaws.associative bind. -Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. - -Search "isMonad" "Store". -Print MonadTypedStore.MonadTypedStore_isMonadTypedStore_mixin. -Print MonadTypedStore.axioms_. -Print isMonadTypedStore.axioms_. -Print isFunctor.axioms_. -Print isMonad.axioms_. - -Let actm (A B : UU0) (f : A -> B) (m : M A) : M B := - mkActo (actm _ _ f (ofActo m)). - -Definition isFunctorTS : isFunctor.axioms_ M. -exists actm. -- move=> A. apply boolp.funext => -[] {A} A m. - by rewrite /actm functor_id. -- move=> A B C g h. apply boolp.funext => m. - case: m h => {A} A m h /=. - by rewrite /actm /= functor_o. -Defined. - -Search "Functor". -Canonical isFunctorTS. - -About Functor. -Print Functor.axioms_. - -Unset Universe Checking. -Canonical Structure FunctorTS : Functor M := Functor.Class isFunctorTS. -Canonical Structure functorTS : functor := Functor.Pack isFunctorTS. -Set Universe Checking. - -Let ret_naturality : naturality [the functor of idfun] M ret. -Proof. by []. Qed. - -Canonical Structure naturalityTSret := isNatural.Build _ _ _ ret_naturality. - -Let ret' : [the functor of idfun] ~> [the functor of M] := - @Nattrans.Pack _ _ ret (Nattrans.Class naturalityTSret). - -Unset Universe Checking. -Let join' : [the functor of M \o M] ~~> [the functor of M] := - fun _ m => bind m idfun. -Set Universe Checking. - -Let actm_bind (a b c : UU0) (f : a -> b) m (g : c -> M a) : - (actm f) (bind m g) = bind m (actm f \o g). -Proof. -congr mkActo. -case: m g => {}c m g /=. -rewrite bindE fmapE 2!bindA /=. -congr (_ >>= _). -Qed. - -Lemma mkActoK A (m : MS _ _ A) : ofActo (mkActo m) = m. -Proof. done. Qed. - -Unset Universe Checking. -Let join'_naturality : naturality _ _ join'. -Proof. -move=> a b h. -unfold join'. -rewrite /=. apply: boolp.funext => mm /=. -unfold hierarchy.actm, isFunctor.actm. -rewrite /= (actm_bind h mm idfun) /= /bind bindA. -apply f_equal. -apply f_equal. -apply boolp.funext => x /=. -rewrite /retS /= fmapE. -case: x h mm => {}a x h mm. -rewrite mkActoK /hierarchy.actm /= /actm /= fmapE !bindretf mkActoK. -reflexivity. -Qed. - -Canonical Structure naturalityTS := isNatural.Build _ _ _ join'_naturality. - -Let join : [the functor of M \o M] ~> [the functor of M] := - @Nattrans.Pack _ _ join' (Nattrans.Class naturalityTS). - - -Canonical Structure isMonadTS : isMonad.axioms_ M isFunctorTS. -exists ret' join bind. -- move=> A B f m. - rewrite /join /= /join' /bind. - apply f_equal. - rewrite bindA. - reflexivity. -- move=> A. - apply boolp.funext => x /=. - case: x => {}A m. - reflexivity. -- move=> A. - apply boolp.funext => x /=. - case: x => {}A m. - rewrite /join' /bind /=. - apply f_equal. - rewrite fmapE /ret. - rewrite bindA /=. - under [fun _ => _]boolp.funext do rewrite bindretf mkActoK. - rewrite bindmret. - reflexivity. -- move=> A /=. - rewrite /join' /bind. - apply boolp.funext => x /=. - apply f_equal. - rewrite fmapE. - rewrite !bindA. - apply f_equal. - apply boolp.funext => y /=. - rewrite bindretf mkActoK. - reflexivity. -Defined. - -Canonical Structure MonadTS : Monad M := Monad.Class isMonadTS. -Canonical Structure monadTS : monad := Monad.Pack isMonadTS. -Set Universe Checking. - -(* Doesn't work -HB.instance Definition xyz := - isMonad_ret_bind.Build M left_neutral right_neutral associative. -*) -(* Since we couldn't build the instance, redefine notations *) -(*Local Notation "m >>= f" := (bind m f). -Local Notation "m >> f" := (bind m (fun=> f)).*) - -(* Basic lemmas for standard library's nth_error *) -Lemma nth_error_set_nth T (def x : T) st n : +Lemma nth_error_set_nth n x : nth_error (set_nth def st n x) n = Some x. Proof. elim: n st => [|z IH] [] // {IH}. @@ -1861,43 +1670,110 @@ elim: z.+1 => [|n <-] //=. by rewrite set_nth_nil. Qed. -Lemma nth_error_rcons_size T (st : seq T) b : +Lemma nth_error_rcons_size b : nth_error (rcons st b) (size st) = Some b. Proof. by elim: st. Qed. -Lemma nth_error_rcons_some T (st : seq T) n a b : +Lemma nth_error_rcons_some n a b : nth_error st n = Some a -> nth_error (rcons st b) n = Some a. Proof. by elim: n st => [|n IH] []. Qed. -Lemma nth_error_set_nth_id A (def : A) st n a : +Lemma nth_error_set_nth_id n a : nth_error st n = Some a -> set_nth def st n a = st. -Proof. elim: n st =>[|n IH] [] //= b st; by [case=> -> | move/IH ->]. Qed. +Proof. elim: n st =>[|n IH] [] //= b st'; by [case=> -> | move/IH ->]. Qed. -Lemma nth_error_set_nth_other A (def : A) st m n a b : +Lemma nth_error_set_nth_other m n a b : m != n -> nth_error st m = Some a -> nth_error (set_nth def st n b) m = Some a. -Proof. elim: m st n => [|m IH] [|c st] [|n] //=; rewrite eqSS => *; exact: IH. +Proof. elim: m st n => [|m IH] [|c st'] [|n] //=; rewrite eqSS => *; exact: IH. Qed. -Lemma nth_error_set_nth_none A (def : A) st m n a b : +Lemma nth_error_set_nth_none m n a b : nth_error st m = None -> nth_error st n = Some a -> nth_error (set_nth def st n b) m = None. -Proof. by elim: m st n => [|m IH] [|c st] [|n] //=; apply IH. Qed. +Proof. by elim: m st n => [|m IH] [|c st'] [|n] //=; apply IH. Qed. -Lemma nth_error_size_set_nth T def (st : seq T) n a b : +Lemma nth_error_size_set_nth n a b : nth_error st n = Some a -> size (set_nth def st n b) = size st. -Proof. by elim: n st => [|n IH] [|c st] //= /IH ->. Qed. +Proof. by elim: n st => [|n IH] [|c st'] //= /IH ->. Qed. -Lemma set_nth_rcons A (def : A) st a b : +Lemma set_nth_rcons a b : set_nth def (rcons st a) (size st) b = rcons st b. -Proof. by elim: st => //= c st ->. Qed. +Proof. by elim: st => //= c st' ->. Qed. -Lemma nth_error_set_nth_rcons T def (st : seq T) n a b c : +Lemma nth_error_set_nth_rcons n a b c : nth_error st n = Some a -> set_nth def (rcons st c) n b = rcons (set_nth def st n b) c. -Proof. by elim: n st => [|n IH] [|d st] //= /IH ->. Qed. +Proof. by elim: n st => [|n IH] [|d st'] //= /IH ->. Qed. +End nth_error. + +Require Import JMeq. + +Module Type MLTYweak. +Parameter ml_type : Set. +Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. +Variant loc : ml_type -> Type := mkloc T : nat -> loc T. +Parameter coq_type0 : ml_type -> Type. +Parameter ml_undef : ml_type. +Parameter undef : coq_type0 ml_undef. +End MLTYweak. + +Module ModelTypedStore (MLtypes : MLTYweak). +Module MLtypes'. +Include MLtypes. +Definition coq_type (M : UU0 -> UU0) := coq_type0. +End MLtypes'. +Module MTypedStore := MonadTypedStore (MLtypes'). +Import MLtypes'. +Import MTypedStore. + +Definition acto (T : UU0) : UU0 := + MS (seq (binding idfun)) [the monad of option_monad] T. +Local Notation M := acto. +(*Local Notation coq_type := (coq_type M).*) +Local Notation coq_type' := (coq_type idfun). + +Definition def : binding idfun := mkbind ml_undef undef. +Local Notation nth_error := List.nth_error. + +Definition cnew T (v : coq_type M T) : M (loc T) := + fun st => let n := size st in + Ret (mkloc T n, rcons st (mkbind T (v : coq_type' T))). + +Definition coerce T1 T2 (v : coq_type M T1) : option (coq_type M T2) := + match ml_type_eq_dec T1 T2 with + | left H => Some (eq_rect _ _ v _ H) + | right _ => None + end. + +Definition cget T (r : loc T) : M (coq_type M T) := + fun st => + if nth_error st (loc_id r) is Some (mkbind T' v) then + if coerce T v is Some u then Ret (u, st) else fail + else fail. + +Definition cput T (r : loc T) (v : coq_type M T) : M unit := + fun st => + let n := loc_id r in + if nth_error st n is Some (mkbind T' _) then + if coerce T' v is Some u then + Ret (tt, set_nth def st n (mkbind T' (u : coq_type' _))) + else fail + else fail. + +Definition cchk T (r : loc T) : M unit := + fun st => + if nth_error st (loc_id r) is Some (mkbind T' u) then + if coerce T u is Some _ then Ret (tt, st) else fail + else fail. + +Definition crun (A : UU0) (m : M A) : option A := + match m nil with + | inl _ => None + | inr (a, _) => Some a + end. (* Make ml_type an eqType *) Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. @@ -1910,29 +1786,26 @@ Definition ml_type_eq_mixin := EqMixin ml_type_eqP. Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. (* Prove the laws *) -Let cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : +Let cnewget T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite nth_error_rcons_size. -rewrite /coerce. +rewrite /cget nth_error_rcons_size /coerce. case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. Qed. -Let cnewgetC T T' (r : loc T) (s : coq_type T') A - (k : loc T' -> coq_type T -> M A) : +Let cnewgetC T T' (r : loc T) (s : coq_type M T') A + (k : loc T' -> coq_type M T -> M A) : cchk r >> (cnew s >>= (fun r' => cget r >>= k r')) = cget r >>= (fun u => cnew s >>= fun r' => k r' u). Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. rewrite {1 3}/coerce. case: ml_type_eq_dec => H /=; last by rewrite bindfailf. @@ -1945,30 +1818,28 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. Qed. -Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : +Let cnewput T (s t : coq_type M T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE bindA. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= [in RHS]bindE /=. +rewrite [in RHS]bindE /= [in RHS]bindE /= /cput. rewrite nth_error_rcons_size. rewrite /coerce. case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. Qed. -Let cnewputC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A +Let cnewputC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A (k : loc T' -> M A) : cchk r >> (cnew s' >>= fun r' => cput r s >> k r') = cput r s >> (cnew s' >>= k). Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cput. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. rewrite {1 3}/coerce. case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. @@ -1985,27 +1856,24 @@ rewrite (nth_error_size_set_nth _ _ Hr). by rewrite (nth_error_set_nth_rcons _ _ _ Hr). Qed. -Let cnewchk T (s : coq_type T) (A : UU0) (k : loc T -> M A) : +Let cnewchk T (s : coq_type M T) (A : UU0) (k : loc T -> M A) : cnew s >>= (fun r => cchk r >> k r) = cnew s >>= k. Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite nth_error_rcons_size. -rewrite /coerce. +rewrite /cchk /cnew nth_error_rcons_size /coerce. by case: ml_type_eq_dec. Qed. -Let cchknewC T1 T2 (r : loc T1) (s : coq_type T2) (A : UU0) +Let cchknewC T1 T2 (r : loc T1) (s : coq_type M T2) (A : UU0) (k : loc T2 -> M A) : cchk r >> (cnew s >>= fun r' => cchk r >> k r') = cchk r >> (cnew s >>= k). Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cnew. case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. case Hc: (coerce T1 u) => [u'|] //. rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /= bindA. @@ -2018,10 +1886,9 @@ Qed. Let cchkC T1 T2 (r1: loc T1) (r2: loc T2) : cchk r1 >> cchk r2 = cchk r2 >> cchk r1. Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk. case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. rewrite bindfailf. case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; @@ -2044,21 +1911,19 @@ Qed. Let cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. Proof. -congr mkActo. apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk. case Hr1 : (nth_error st (loc_id r)) => [[T1' v1]|] //=. case Hc: (coerce T v1) => [u1|] //. by rewrite bindE /= bindE /= Hr1 Hc. Qed. -Let cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type T2 -> M A) : +Let cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type M T2 -> M A) : cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. rewrite bindfailf. case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; @@ -2081,39 +1946,36 @@ do! rewrite bindE /=. by rewrite /bindS /= MS_mapE /= fmapE Hr1 Hc. Qed. -Let cchkget T (r : loc T) (A: UU0) (k : coq_type T -> M A) : +Let cchkget T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : cchk r >> (cget r >>= k) = cget r >>= k. Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. case Hc: (coerce T u) => [u'|] //. do! rewrite bindE /=. by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. Qed. -Let cgetchk T (r : loc T) (A: UU0) (k : coq_type T -> M A) : +Let cgetchk T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : cget r >>= (fun s => cchk r >> k s) = cget r >>= k. Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cchk. case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. case Hc: (coerce T u) => [u'|] //. do! rewrite bindE /=. by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. Qed. -Let cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : +Let cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : cchk r1 >> cput r2 s = cput r2 s >> cchk r1. Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cput. case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. rewrite bindfailf. case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; @@ -2153,12 +2015,11 @@ rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. by case: ml_type_eq_dec. Qed. -Let cchkput T (r : loc T) (s : coq_type T) : +Let cchkput T (r : loc T) (s : coq_type M T) : cchk r >> cput r s = cput r s. Proof. -congr mkActo. apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cput. case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. rewrite /coerce. case: ml_type_eq_dec => HT; last by case: ml_type_eq_dec => // /esym //. @@ -2171,12 +2032,11 @@ case: ml_type_eq_dec => // HT. by rewrite -eq_rect_eq. Qed. -Let cputchk T (r : loc T) (s : coq_type T) : +Let cputchk T (r : loc T) (s : coq_type M T) : cput r s >> cchk r = cput r s. Proof. -congr mkActo. apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cchk. case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. rewrite /coerce. case: ml_type_eq_dec => HT //. @@ -2187,13 +2047,12 @@ rewrite nth_error_set_nth. by case: ml_type_eq_dec. Qed. -Let cputput T (r : loc T) (s s' : coq_type T) : +Let cputput T (r : loc T) (s s' : coq_type M T) : cput r s >> cput r s' = cput r s'. Proof. -congr mkActo. apply/boolp.funext => st. case: r s s' => {}T n s s' /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. rewrite {1 3}/coerce. case: ml_type_eq_dec => H /=; last by rewrite bindfailf. @@ -2204,14 +2063,13 @@ rewrite -eq_rect_eq. by rewrite set_set_nth eqxx. Qed. -Let cputget T (r : loc T) (s : coq_type T) (A : UU0) (k : coq_type T -> M A) : +Let cputget T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. -congr mkActo. apply/boolp.funext => st /=. case: r s k => {}T n s k /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. rewrite {1 3}/coerce. case: ml_type_eq_dec => H /=; last by rewrite bindfailf. @@ -2226,9 +2084,8 @@ Qed. Let cgetputchk T (r : loc T) : cget r >>= cput r = cchk r. Proof. -congr mkActo. apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput /cchk. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. rewrite /coerce. case: (ml_type_eq_dec T1 T) => H //. @@ -2238,13 +2095,12 @@ case: (ml_type_eq_dec T T) => H //. by rewrite -eq_rect_eq nth_error_set_nth_id. Qed. -Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : +Let cgetget T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. rewrite /coerce. case: (ml_type_eq_dec T1 T) => H //. @@ -2257,14 +2113,13 @@ by rewrite -eq_rect_eq. Qed. Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) - (k : coq_type T1 -> coq_type T2 -> M A) : + (k : coq_type M T1 -> coq_type M T2 -> M A) : cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). Proof. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. case Hr1: (nth_error _ _) => [[T1' u]|]; last first. rewrite bindfailf. case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. @@ -2299,16 +2154,15 @@ case: (ml_type_eq_dec T1 T1) => // H. by rewrite -eq_rect_eq. Qed. -Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) - (s2 : coq_type T2) (A : UU0) : +Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (s2 : coq_type M T2) (A : UU0) : loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. Proof. move=> H. -congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. case Hr1: (nth_error _ _) => [[T1' u]|]; last first. rewrite bindfailf. case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. @@ -2367,17 +2221,16 @@ case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq set_set_nth (negbTE Hr). Qed. -Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) - (A : UU0) (k : coq_type T2 -> M A) : +Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (A : UU0) (k : coq_type M T2 -> M A) : loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). Proof. move=> Hr. -congr mkActo. apply/boolp.funext => st /=. rewrite !bindA. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. case Hr1: (nth_error _ _) => [[T1' u]|]; last first. rewrite bindfailf. case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. @@ -2409,9 +2262,8 @@ by rewrite -eq_rect_eq bindE. Qed. Let crunret (A B : UU0) (m : M A) (s : B) : - crun m -> crun (m >> ret s) = Some s. + crun m -> crun (m >> Ret s) = Some s. Proof. -case: m => T /= m. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. by case Hm: (m [::]). Qed. @@ -2419,17 +2271,15 @@ Qed. Let crunskip : crun skip = Some tt. Proof. by []. Qed. -Let crunnew (A : UU0) T (m : M A) (s : coq_type T) : +Let crunnew (A : UU0) T (m : M A) (s : coq_type M T) : crun m -> crun (m >> cnew s). Proof. -case: m => U /= m. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. by case Hm: (m [::]). Qed. -Unset Universe Checking. Definition isMonadTypedStoreModel : - isMonadTypedStore.axioms_ M isFunctorTS isMonadTS. + isMonadTypedStore.axioms_ M (HB_unnamed_mixin_21 _ _) (HB_unnamed_mixin_22 _ _). Proof. exists cnew cget cput cchk crun. - exact cnewget. @@ -2457,7 +2307,14 @@ exists cnew cget cput cchk crun. - exact crunskip. - exact crunnew. Defined. -Set Universe Checking. + +(* +HB.instance Definition _ := + isMonadTypedStore.Build M cnewget cnewgetC cnewput cnewputC cputput + cputget cgetputchk cgetget cgetC cputC cputgetC cnewchk cchknewC + cchkgetC cchkget cgetchk cchkputC cchkput cputchk cchkC cchkdup + crunret crunskip crunnew. +*) End ModelTypedStore. (* TODO? diff --git a/typed_store_model.v b/typed_store_model.v new file mode 100644 index 00000000..43f09d4a --- /dev/null +++ b/typed_store_model.v @@ -0,0 +1,762 @@ +From mathcomp Require Import all_ssreflect. +From mathcomp Require Import finmap. +From mathcomp Require boolp. +From mathcomp Require Import classical_sets. +From infotheo Require convex classical_sets_ext. +Require Import monae_lib. +From HB Require Import structures. +Require Import hierarchy monad_lib fail_lib state_lib trace_lib. +Require Import monad_transformer monad_model. +Require Import JMeq. + +(******************************************************************************) +(* Models for typed store *) +(* (separate file as it requires disabling various sanity checks) *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope monae_scope. + +Module Type MLTYdef. +Include MLTY. +Parameter ml_undef : ml_type. +Parameter undef : forall M, coq_type M ml_undef. +End MLTYdef. + +Module ModelTypedStore (MLtypes : MLTYdef). +Module MTypedStore := MonadTypedStore (MLtypes). +Import MLtypes. +Import MTypedStore. + +#[bypass_check(positivity)] +Inductive acto : UU0 -> UU0 := + mkActo : forall T : UU0, + MS (seq (binding acto)) [the monad of option_monad] T -> acto T. +Local Notation M := acto. +Local Notation coq_type := (coq_type M). + +Definition ofActo T (m : M T) + : MS (seq (binding M)) [the monad of option_monad] T := + let: mkActo _ m' := m in m'. + +Definition def := mkbind ml_undef (undef M). +Local Notation nth_error := List.nth_error. + +Definition cnew T (v : coq_type T) : M (loc T) := + mkActo (fun st => let n := size st in Ret (mkloc T n, rcons st (mkbind T v))). + +Definition coerce (T1 T2 : ml_type) (v : coq_type T1) : option (coq_type T2) := + match ml_type_eq_dec T1 T2 with + | left H => Some (eq_rect _ _ v _ H) + | right _ => None + end. + +Definition cget T (r : loc T) : M (coq_type T) := + mkActo (fun st => + if nth_error st (loc_id r) is Some (mkbind T' v) then + if coerce T v is Some u then Ret (u, st) else fail + else fail). + +Definition cput T (r : loc T) (v : coq_type T) : M unit := + mkActo (fun st => + let n := loc_id r in + if nth_error st n is Some (mkbind T' _) then + if coerce T' v is Some u then + Ret (tt, set_nth def st n (mkbind T' u)) + else fail + else fail). + +Definition cchk T (r : loc T) : M unit := + mkActo (fun st => + if nth_error st (loc_id r) is Some (mkbind T' u) then + if coerce T u is Some _ then Ret (tt, st) else fail + else fail). + +Definition crun (A : UU0) (m : M A) : option A := + match ofActo m nil with + | inl _ => None + | inr (a, _) => Some a + end. + +Let ret : forall A, idfun A -> M A := fun A a => mkActo (Ret a). +Let bind A B (m : M A) (f : A -> M B) : M B := + mkActo (ofActo m >>= (fun a => ofActo (f a))). +Let left_neutral : BindLaws.left_neutral bind ret. +Proof. by move=> A B a f; rewrite /bind /ret bindretf; case: (f a). Qed. +Let right_neutral : BindLaws.right_neutral bind ret. +Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. +Let associative : BindLaws.associative bind. +Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. + +Let actm (A B : UU0) (f : A -> B) (m : M A) : M B := + mkActo (actm _ _ f (ofActo m)). + +Definition isFunctorTS : isFunctor.axioms_ M. +exists actm. +- move=> A. apply boolp.funext => -[] {A} A m. + by rewrite /actm functor_id. +- move=> A B C g h. apply boolp.funext => m. + case: m h => {A} A m h /=. + by rewrite /actm /= functor_o. +Defined. + +Unset Universe Checking. +Canonical Structure FunctorTS : Functor M := Functor.Class isFunctorTS. +Canonical Structure functorTS : functor := Functor.Pack isFunctorTS. +Set Universe Checking. + +Let ret_naturality : naturality [the functor of idfun] M ret. +Proof. by []. Qed. + +Canonical Structure naturalityTSret := isNatural.Build _ _ _ ret_naturality. + +Let ret' : [the functor of idfun] ~> [the functor of M] := + @Nattrans.Pack _ _ ret (Nattrans.Class naturalityTSret). + +Unset Universe Checking. +Let join' : [the functor of M \o M] ~~> [the functor of M] := + fun _ m => bind m idfun. +Set Universe Checking. + +Let actm_bind (a b c : UU0) (f : a -> b) m (g : c -> M a) : + (actm f) (bind m g) = bind m (actm f \o g). +Proof. +congr mkActo. +case: m g => {}c m g /=. +rewrite bindE fmapE 2!bindA /=. +reflexivity. +Qed. + +Lemma mkActoK A (m : MS _ _ A) : ofActo (mkActo m) = m. +Proof. done. Qed. + +Unset Universe Checking. +Let join'_naturality : naturality _ _ join'. +Proof. +move=> a b h. +unfold join'. +rewrite /=. apply: boolp.funext => mm /=. +unfold hierarchy.actm, isFunctor.actm. +rewrite /= (actm_bind h mm idfun) /= /bind bindA. +apply f_equal. +apply f_equal. +apply boolp.funext => x /=. +rewrite /retS /= fmapE. +case: x h mm => {}a x h mm. +rewrite mkActoK /hierarchy.actm /= /actm /= fmapE !bindretf mkActoK. +reflexivity. +Qed. + +Canonical Structure naturalityTS := isNatural.Build _ _ _ join'_naturality. + +Let join : [the functor of M \o M] ~> [the functor of M] := + @Nattrans.Pack _ _ join' (Nattrans.Class naturalityTS). + +Canonical Structure isMonadTS : isMonad.axioms_ M isFunctorTS. +exists ret' join bind. +- move=> A B f m. + rewrite /join /= /join' /bind. + apply f_equal. + rewrite bindA. + reflexivity. +- move=> A. + apply boolp.funext => x /=. + case: x => {}A m. + reflexivity. +- move=> A. + apply boolp.funext => x /=. + case: x => {}A m. + rewrite /join' /bind /=. + apply f_equal. + rewrite fmapE /ret. + rewrite bindA /=. + under [fun _ => _]boolp.funext do rewrite bindretf mkActoK. + rewrite bindmret. + reflexivity. +- move=> A /=. + rewrite /join' /bind. + apply boolp.funext => x /=. + apply f_equal. + rewrite fmapE. + rewrite !bindA. + apply f_equal. + apply boolp.funext => y /=. + rewrite bindretf mkActoK. + reflexivity. +Defined. + +Canonical Structure MonadTS : Monad M := Monad.Class isMonadTS. +Canonical Structure monadTS : monad := Monad.Pack isMonadTS. +Set Universe Checking. + +(* Doesn't work +HB.instance Definition xyz := + isMonad_ret_bind.Build M left_neutral right_neutral associative. +*) +(* Since we couldn't build the instance, redefine notations *) +(*Local Notation "m >>= f" := (bind m f). +Local Notation "m >> f" := (bind m (fun=> f)).*) + +(* Make ml_type an eqType *) +Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. +Lemma ml_type_eqP : Equality.axiom ml_type_eqb. +Proof. +rewrite /ml_type_eqb => T1 T2. +by case: ml_type_eq_dec; constructor. +Qed. +Definition ml_type_eq_mixin := EqMixin ml_type_eqP. +Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. + +(* Prove the laws *) +Let cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : + cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite nth_error_rcons_size. +rewrite /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cnewgetC T T' (r : loc T) (s : coq_type T') A + (k : loc T' -> coq_type T -> M A) : + cchk r >> (cnew s >>= (fun r' => cget r >>= k r')) = + cget r >>= (fun u => cnew s >>= fun r' => k r' u). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T1. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite (nth_error_rcons_some _ Hr). +rewrite /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : + cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE bindA. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= [in RHS]bindE /=. +rewrite nth_error_rcons_size. +rewrite /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. +Qed. + +Let cnewputC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A + (k : loc T' -> M A) : + cchk r >> (cnew s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew s' >>= k). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. +subst T1. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite (nth_error_rcons_some _ Hr). +rewrite /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /=. +rewrite (nth_error_size_set_nth _ _ Hr). +by rewrite (nth_error_set_nth_rcons _ _ _ Hr). +Qed. + +Let cnewchk T (s : coq_type T) (A : UU0) (k : loc T -> M A) : + cnew s >>= (fun r => cchk r >> k r) = cnew s >>= k. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite nth_error_rcons_size. +rewrite /coerce. +by case: ml_type_eq_dec. +Qed. + +Let cchknewC T1 T2 (r : loc T1) (s : coq_type T2) (A : UU0) + (k : loc T2 -> M A) : + cchk r >> (cnew s >>= fun r' => cchk r >> k r') = cchk r >> (cnew s >>= k). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +case Hc: (coerce T1 u) => [u'|] //. +rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /= bindA. +rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /= bindA. +rewrite (nth_error_rcons_some _ Hr) Hc. +do! rewrite bindE /=. +by rewrite /bindS /= MS_mapE /= fmapE bindA. +Qed. + +Let cchkC T1 T2 (r1: loc T1) (r2: loc T2) : + cchk r1 >> cchk r2 = cchk r2 >> cchk r1. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. + rewrite bindfailf. + case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; + last by rewrite bindfailf. + case Hc: (coerce T2 v2) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= Hr1. +case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. + rewrite bindfailf. + case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= Hr2. +case Hc: (coerce T1 v1) => [u1|]; last first. + rewrite bindfailf. + case Hc2: (coerce T2 v2) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= Hr1 Hc. +rewrite bindE /= bindE /= Hr2. +case Hc2: (coerce T2 v2) => [u2|]; last by rewrite !bindfailf. +do! rewrite bindE /=. +by rewrite Hr1 Hc. +Qed. + +Let cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1 : (nth_error st (loc_id r)) => [[T1' v1]|] //=. +case Hc: (coerce T v1) => [u1|] //. +by rewrite bindE /= bindE /= Hr1 Hc. +Qed. + +Let cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type T2 -> M A) : + cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. + rewrite bindfailf. + case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; + last by rewrite bindfailf. + case Hc: (coerce T2 v2) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1 bindfailf. +case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. + rewrite bindfailf. + case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr2 bindfailf. +case Hc: (coerce T1 v1) => [u1|]; last first. + rewrite bindfailf. + case Hc2: (coerce T2 v2) => [u|]; last by rewrite bindfailf. + rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1 Hc. + by rewrite bindfailf. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite Hr2. +case Hc2: (coerce T2 v2) => [u2|]; last by rewrite !bindfailf. +do! rewrite bindE /=. +by rewrite /bindS /= MS_mapE /= fmapE Hr1 Hc. +Qed. + +Let cchkget T (r : loc T) (A: UU0) (k : coq_type T -> M A) : + cchk r >> (cget r >>= k) = cget r >>= k. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +case Hc: (coerce T u) => [u'|] //. +do! rewrite bindE /=. +by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. +Qed. + +Let cgetchk T (r : loc T) (A: UU0) (k : coq_type T -> M A) : + cget r >>= (fun s => cchk r >> k s) = cget r >>= k. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +case Hc: (coerce T u) => [u'|] //. +do! rewrite bindE /=. +by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. +Qed. + +Let cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : + cchk r1 >> cput r2 s = cput r2 s >> cchk r1. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. + rewrite bindfailf. + case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; + last by rewrite bindfailf. + case Hc: (coerce T2' s) => [u|]; last by rewrite bindfailf. + rewrite bindE /= /bindS /= bindE /=. + by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). +case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. + rewrite bindfailf. + case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= Hr2. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec T1' T1) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec T2 T2') => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE/= bindE/=. + case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. + rewrite -Hr nth_error_set_nth /coerce. + move: Hr2; rewrite -Hr Hr1 => -[] HT2. + subst T1'. + by case: ml_type_eq_dec. + rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. + by case: ml_type_eq_dec. +rewrite bindE /= bindE /= Hr2 {1}/coerce. +case: ml_type_eq_dec => HT2 //. +subst T1' T2'. +rewrite -!eq_rect_eq. +rewrite bindE /= bindE /=. +case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. + rewrite -Hr nth_error_set_nth /coerce. + move: Hr2; rewrite -Hr Hr1 => -[] HT2. + subst T2. + by case: ml_type_eq_dec. +rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. +by case: ml_type_eq_dec. +Qed. + +Let cchkput T (r : loc T) (s : coq_type T) : + cchk r >> cput r s = cput r s. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +rewrite /coerce. +case: ml_type_eq_dec => HT; last by case: ml_type_eq_dec => // /esym //. +subst T'. +case: ml_type_eq_dec => // HT. +rewrite -eq_rect_eq {HT}. +do! rewrite bindE /=. +rewrite Hr. +case: ml_type_eq_dec => // HT. +by rewrite -eq_rect_eq. +Qed. + +Let cputchk T (r : loc T) (s : coq_type T) : + cput r s >> cchk r = cput r s. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +rewrite /coerce. +case: ml_type_eq_dec => HT //. +subst T'. +rewrite -!eq_rect_eq. +do! rewrite bindE /=. +rewrite nth_error_set_nth. +by case: ml_type_eq_dec. +Qed. + +Let cputput T (r : loc T) (s s' : coq_type T) : + cput r s >> cput r s' = cput r s'. +Proof. +congr mkActo. +apply/boolp.funext => st. +case: r s s' => {}T n s s' /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T'. +rewrite !bindretf /= nth_error_set_nth /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +by rewrite set_set_nth eqxx. +Qed. + +Let cputget T (r : loc T) (s : coq_type T) (A : UU0) (k : coq_type T -> M A) : + cput r s >> (cget r >>= k) = cput r s >> k s. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +case: r s k => {}T n s k /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T'. +rewrite !bindretf /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite nth_error_set_nth. +rewrite /coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +by rewrite -eq_rect_eq !bindretf. +Qed. + +Let cgetputchk T (r : loc T) : cget r >>= cput r = cchk r. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite /coerce. +case: (ml_type_eq_dec T1 T) => H //. +subst T1. +rewrite bindE /= bindE /= Hr. +case: (ml_type_eq_dec T T) => H //. +by rewrite -eq_rect_eq nth_error_set_nth_id. +Qed. + +Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : + cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite /coerce. +case: (ml_type_eq_dec T1 T) => H //. +subst T1. +rewrite bindE /= bindE /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite Hr. +case: (ml_type_eq_dec T T) => H //. +by rewrite -eq_rect_eq. +Qed. + +Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) + (k : coq_type T1 -> coq_type T2 -> M A) : + cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = + cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1: (nth_error _ _) => [[T1' u]|]; last first. + rewrite bindfailf. + case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. + case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. + rewrite !bindE /=. + by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. +case Hr2: (nth_error _ _) => [[T2' v]|]; last first. + rewrite bindfailf. + case Hc: (coerce T1 u) => [v|]; last by rewrite bindfailf. + rewrite !bindE /=. + by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr2 bindfailf. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec _ _) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE /=. + rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr1 {1}/coerce. + case: (ml_type_eq_dec _ _) => //; by rewrite bindfailf. +subst T1'. +rewrite -eq_rect_eq. +rewrite bindE /=. +rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr2 {1}/coerce. +case: (ml_type_eq_dec T2' T2) => HT2; last by rewrite !bindfailf. +subst T2'. +rewrite -eq_rect_eq. +rewrite !bindE /=. +rewrite !bindE /=. +rewrite /bindS MS_mapE /= fmapE /= bindA /= Hr1 /coerce. +case: (ml_type_eq_dec T1 T1) => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) + (s2 : coq_type T2) (A : UU0) : + loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> + cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. +Proof. +move=> H. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1: (nth_error _ _) => [[T1' u]|]; last first. + rewrite bindfailf. + case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. + case Hc: (coerce T2' s2) => [u|]; last by rewrite bindfailf. + rewrite !bindE /=. + rewrite !bindE /=. + case/boolP: (loc_id r1 == loc_id r2) => Hr. + by rewrite -(eqP Hr) Hr1 in Hr2. + by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). +case/boolP: (loc_id r1 == loc_id r2) H => [/eqP|] Hr /= H. + rewrite -Hr Hr1. + case/boolP: (T1 == T2) => /eqP HT; last first. + rewrite /coerce. + case: (ml_type_eq_dec T1 T1') => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec T2 T1') => HT2; last by rewrite bindfailf. + subst T1'. + rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. + by case: ml_type_eq_dec. + subst T1'. + rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. + by case: ml_type_eq_dec => [/esym|] //. + subst T2. + rewrite /coerce. + case: ml_type_eq_dec => HT1 //. + subst T1'. + rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /= !nth_error_set_nth. + case: ml_type_eq_dec => HT1 //. + rewrite -!eq_rect_eq. + case: H => // H. + by move: (JMeq_eq H) => <-. +case Hr2: (nth_error _ _) => [[T2' v]|]; last first. + rewrite bindfailf. + case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. + rewrite !bindE /=. + rewrite !bindE /=. + by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1). +rewrite {1 3}/coerce. +case: (ml_type_eq_dec _ _) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE /=. + rewrite !bindE /=. + rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. + by case: ml_type_eq_dec. +subst T1'. +rewrite -eq_rect_eq /coerce. +rewrite bindE /= bindE /= (nth_error_set_nth_other _ _ _ Hr2); + last by rewrite eq_sym. +case: ml_type_eq_dec => HT2' //. +subst T2'. +rewrite -eq_rect_eq bindE /= bindE /= (nth_error_set_nth_other _ _ Hr Hr1). +case: ml_type_eq_dec => // HT1. +by rewrite -eq_rect_eq set_set_nth (negbTE Hr). +Qed. + +Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) + (A : UU0) (k : coq_type T2 -> M A) : + loc_id r1 != loc_id r2 -> + cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). +Proof. +move=> Hr. +congr mkActo. +apply/boolp.funext => st /=. +rewrite !bindA. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hr1: (nth_error _ _) => [[T1' u]|]; last first. + rewrite bindfailf. + case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. + case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. + by rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. +case Hr2: (nth_error _ _) => [[T2' v]|]; last first. + rewrite bindfailf. + case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. + rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE. + by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1) bindfailf. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec _ _) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE Hr1 /coerce. + by case: ml_type_eq_dec. +subst T1'. +rewrite -eq_rect_eq /coerce. +rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite (nth_error_set_nth_other _ _ _ Hr2); last by rewrite eq_sym. +case: ml_type_eq_dec => HT2' //. +subst T2'. +rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /=. +rewrite /bindS /= MS_mapE /= fmapE Hr1. +case: ml_type_eq_dec => // HT1. +by rewrite -eq_rect_eq bindE. +Qed. + +Let crunret (A B : UU0) (m : M A) (s : B) : + crun m -> crun (m >> ret s) = Some s. +Proof. +case: m => T /= m. +rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +by case Hm: (m [::]). +Qed. + +Let crunskip : crun skip = Some tt. +Proof. by []. Qed. + +Let crunnew (A : UU0) T (m : M A) (s : coq_type T) : + crun m -> crun (m >> cnew s). +Proof. +case: m => U /= m. +rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +by case Hm: (m [::]). +Qed. + +Unset Universe Checking. +Definition isMonadTypedStoreModel : + isMonadTypedStore.axioms_ M isFunctorTS isMonadTS. +Proof. +exists cnew cget cput cchk crun. +- exact cnewget. +- exact cnewgetC. +- exact cnewput. +- exact cnewputC. +- exact cputput. +- exact cputget. +- exact cgetputchk. +- exact cgetget. +- exact cgetC. +- exact cputC. +- exact cputgetC. +- exact cnewchk. +- exact cchknewC. +- exact cchkgetC. +- exact cchkget. +- exact cgetchk. +- exact cchkputC. +- exact cchkput. +- exact cputchk. +- exact cchkC. +- exact cchkdup. +- exact crunret. +- exact crunskip. +- exact crunnew. +Defined. +Set Universe Checking. +End ModelTypedStore. From 1049c4a191369d84e042ff231cba8cf94cd91f0a Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 31 Jan 2023 15:01:48 +0900 Subject: [PATCH 12/82] make definition closer to expected --- monad_model.v | 38 +++++++------------------------------- 1 file changed, 7 insertions(+), 31 deletions(-) diff --git a/monad_model.v b/monad_model.v index 32d4ced3..5b4828c7 100644 --- a/monad_model.v +++ b/monad_model.v @@ -2278,37 +2278,13 @@ rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. by case Hm: (m [::]). Qed. -Definition isMonadTypedStoreModel : - isMonadTypedStore.axioms_ M (HB_unnamed_mixin_21 _ _) (HB_unnamed_mixin_22 _ _). -Proof. -exists cnew cget cput cchk crun. -- exact cnewget. -- exact cnewgetC. -- exact cnewput. -- exact cnewputC. -- exact cputput. -- exact cputget. -- exact cgetputchk. -- exact cgetget. -- exact cgetC. -- exact cputC. -- exact cputgetC. -- exact cnewchk. -- exact cchknewC. -- exact cchkgetC. -- exact cchkget. -- exact cgetchk. -- exact cchkputC. -- exact cchkput. -- exact cputchk. -- exact cchkC. -- exact cchkdup. -- exact crunret. -- exact crunskip. -- exact crunnew. -Defined. - -(* +Canonical Structure isMonadTypedStoreModel := + isMonadTypedStore.Build M cnewget cnewgetC cnewput cnewputC cputput + cputget cgetputchk cgetget cgetC cputC cputgetC cnewchk cchknewC + cchkgetC cchkget cgetchk cchkputC cchkput cputchk cchkC cchkdup + crunret crunskip crunnew. + +(* Fails HB.instance Definition _ := isMonadTypedStore.Build M cnewget cnewgetC cnewput cnewputC cputput cputget cgetputchk cgetget cgetC cputC cputgetC cnewchk cchknewC From 17e836f818017a1ceb9c5a61e53f250fd1db8fbd Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Wed, 1 Feb 2023 15:06:49 +0900 Subject: [PATCH 13/82] remove hint incompatible disabling universe checking --- typed_store_model.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 43f09d4a..2a297726 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -1,8 +1,7 @@ From mathcomp Require Import all_ssreflect. From mathcomp Require Import finmap. From mathcomp Require boolp. -From mathcomp Require Import classical_sets. -From infotheo Require convex classical_sets_ext. +#[local] Remove Hints boolp.Prop_irrelevance : core. Require Import monae_lib. From HB Require Import structures. Require Import hierarchy monad_lib fail_lib state_lib trace_lib. From 3d4ad81619ebf6c8fa94ec44c3f9df49f334cbd2 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Wed, 1 Feb 2023 19:55:05 +0900 Subject: [PATCH 14/82] use specialized version of funext --- example_typed_store.v | 35 +++++++++++++++++------------------ 1 file changed, 17 insertions(+), 18 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 706a7c01..c21c9960 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -72,6 +72,9 @@ Proof. by apply boolp.funext => r; rewrite bindmret. Qed. Lemma crunnew0 (M : typedStoreMonad) T s : crun (cnew T s : M _). Proof. by rewrite -[cnew _ _]bindskipf crunnew // crunskip. Qed. +Lemma funext {T U : UU0} [f g : T -> U] : f =1 g -> (fun x => f x) = g. +Proof. exact: boolp.funext. Qed. + Section factorial. Variable M : typedStoreMonad. Notation coq_type := (@MLtypes.coq_type M). @@ -89,11 +92,11 @@ set s := 1. have smn : s * fact_rec m = fn by rewrite mul1n. elim: m s smn => [|m IH] s /= smn. rewrite /fact_ref -smn muln1. - under [fun r => _]boolp.funext do rewrite bindretf. + under funext do rewrite bindretf. by rewrite cgetret cnewget crunret // crunnew0. -under [fun r => _]boolp.funext do rewrite bindA. +under funext do rewrite bindA. rewrite cnewget. -under [fun r => _]boolp.funext do rewrite bindA. +under funext do rewrite bindA. by rewrite cnewput IH // (mulnC m.+1) -mulnA. Qed. End factorial. @@ -122,18 +125,17 @@ set x := 1. pose y := x. rewrite -{2}/y. pose i := n. -rewrite -{1}/i. +rewrite -[in LHS]/i. have : (x, y) = (fibo_rec (n-i), fibo_rec (n.+1-i)). by rewrite subnn -addn1 addKn. have : i <= n by []. elim: i x y => [|i IH] x y Hi. rewrite !subn0 => -[] -> ->. rewrite -/(fibo_rec n.+1). - under [fun a => _]boolp.funext do - under [fun b => _]boolp.funext do rewrite [fibo_ref _ _ _]/= bindskipf. - under [fun a => _]boolp.funext do rewrite cgetret. + under funext do under funext do rewrite [fibo_ref _ _ _]/= bindskipf. + under funext do rewrite cgetret. rewrite -cnewchk. - under [fun b => _]boolp.funext do rewrite cnewgetC. + under funext do rewrite cnewgetC. by rewrite cnewget -bindA crunret // crunnew // crunnew0. rewrite subSS => -[] Hx Hy. rewrite -(IH y (x + y) (ltnW Hi)); last first. @@ -144,19 +146,16 @@ rewrite -(IH y (x + y) (ltnW Hi)); last first. rewrite subSS. by rewrite -addn2 -addn1 -addnBAC // -addnBAC // addn2 addn1. rewrite /=. -under [fun a => _]boolp.funext do -under [fun a => _]boolp.funext do rewrite !bindA. +under funext do under funext do rewrite !bindA. rewrite -cnewchk. -under [fun b => _]boolp.funext do rewrite cnewgetC. +under funext do rewrite cnewgetC. rewrite cnewget. -under [fun a => _]boolp.funext do -under [fun a => _]boolp.funext do rewrite !bindA. -under [fun a => _]boolp.funext do rewrite cnewget. -under [fun a => _]boolp.funext do -under [fun a => _]boolp.funext do rewrite !bindA. +under funext do under funext do rewrite !bindA. +under funext do rewrite cnewget. +under funext do under funext do rewrite !bindA. rewrite -cnewchk. -under [fun a => _]boolp.funext do rewrite cnewputC. +under funext do rewrite cnewputC. rewrite cnewput. -by under [fun a => _]boolp.funext do rewrite cnewput. +by under funext do rewrite cnewput. Qed. End fibonacci. From 2e71b3dbae9842a4d3760cbb14d949c8784dcb71 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Wed, 1 Feb 2023 23:11:23 +0900 Subject: [PATCH 15/82] add cchknew to infer reference inequalities --- typed_store_model.v | 91 +++++++++++++++++++++++++++------------------ 1 file changed, 55 insertions(+), 36 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 2a297726..6ea32762 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -80,17 +80,17 @@ Definition crun (A : UU0) (m : M A) : option A := | inr (a, _) => Some a end. -Let ret : forall A, idfun A -> M A := fun A a => mkActo (Ret a). -Let bind A B (m : M A) (f : A -> M B) : M B := +Definition ret : forall A, idfun A -> M A := fun A a => mkActo (Ret a). +Definition bind A B (m : M A) (f : A -> M B) : M B := mkActo (ofActo m >>= (fun a => ofActo (f a))). -Let left_neutral : BindLaws.left_neutral bind ret. +Definition left_neutral : BindLaws.left_neutral bind ret. Proof. by move=> A B a f; rewrite /bind /ret bindretf; case: (f a). Qed. -Let right_neutral : BindLaws.right_neutral bind ret. +Definition right_neutral : BindLaws.right_neutral bind ret. Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. -Let associative : BindLaws.associative bind. +Definition associative : BindLaws.associative bind. Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. -Let actm (A B : UU0) (f : A -> B) (m : M A) : M B := +Definition actm (A B : UU0) (f : A -> B) (m : M A) : M B := mkActo (actm _ _ f (ofActo m)). Definition isFunctorTS : isFunctor.axioms_ M. @@ -107,20 +107,20 @@ Canonical Structure FunctorTS : Functor M := Functor.Class isFunctorTS. Canonical Structure functorTS : functor := Functor.Pack isFunctorTS. Set Universe Checking. -Let ret_naturality : naturality [the functor of idfun] M ret. +Definition ret_naturality : naturality [the functor of idfun] M ret. Proof. by []. Qed. Canonical Structure naturalityTSret := isNatural.Build _ _ _ ret_naturality. -Let ret' : [the functor of idfun] ~> [the functor of M] := +Definition ret' : [the functor of idfun] ~> [the functor of M] := @Nattrans.Pack _ _ ret (Nattrans.Class naturalityTSret). Unset Universe Checking. -Let join' : [the functor of M \o M] ~~> [the functor of M] := +Definition join' : [the functor of M \o M] ~~> [the functor of M] := fun _ m => bind m idfun. Set Universe Checking. -Let actm_bind (a b c : UU0) (f : a -> b) m (g : c -> M a) : +Definition actm_bind (a b c : UU0) (f : a -> b) m (g : c -> M a) : (actm f) (bind m g) = bind m (actm f \o g). Proof. congr mkActo. @@ -133,7 +133,7 @@ Lemma mkActoK A (m : MS _ _ A) : ofActo (mkActo m) = m. Proof. done. Qed. Unset Universe Checking. -Let join'_naturality : naturality _ _ join'. +Definition join'_naturality : naturality _ _ join'. Proof. move=> a b h. unfold join'. @@ -151,7 +151,7 @@ Qed. Canonical Structure naturalityTS := isNatural.Build _ _ _ join'_naturality. -Let join : [the functor of M \o M] ~> [the functor of M] := +Definition join : [the functor of M \o M] ~> [the functor of M] := @Nattrans.Pack _ _ join' (Nattrans.Class naturalityTS). Canonical Structure isMonadTS : isMonad.axioms_ M isFunctorTS. @@ -210,7 +210,7 @@ Definition ml_type_eq_mixin := EqMixin ml_type_eqP. Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. (* Prove the laws *) -Let cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : +Definition cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. congr mkActo. @@ -224,7 +224,7 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. Qed. -Let cnewgetC T T' (r : loc T) (s : coq_type T') A +Definition cnewgetC T T' (r : loc T) (s : coq_type T') A (k : loc T' -> coq_type T -> M A) : cchk r >> (cnew s >>= (fun r' => cget r >>= k r')) = cget r >>= (fun u => cnew s >>= fun r' => k r' u). @@ -245,7 +245,7 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. Qed. -Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : +Definition cnewput T (s t : coq_type T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. congr mkActo. @@ -260,7 +260,7 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. Qed. -Let cnewputC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A +Definition cnewputC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A (k : loc T' -> M A) : cchk r >> (cnew s' >>= fun r' => cput r s >> k r') = cput r s >> (cnew s' >>= k). @@ -285,7 +285,7 @@ rewrite (nth_error_size_set_nth _ _ Hr). by rewrite (nth_error_set_nth_rcons _ _ _ Hr). Qed. -Let cnewchk T (s : coq_type T) (A : UU0) (k : loc T -> M A) : +Definition cnewchk T (s : coq_type T) (A : UU0) (k : loc T -> M A) : cnew s >>= (fun r => cchk r >> k r) = cnew s >>= k. Proof. congr mkActo. @@ -298,7 +298,7 @@ rewrite /coerce. by case: ml_type_eq_dec. Qed. -Let cchknewC T1 T2 (r : loc T1) (s : coq_type T2) (A : UU0) +Definition cchknewC T1 T2 (r : loc T1) (s : coq_type T2) (A : UU0) (k : loc T2 -> M A) : cchk r >> (cnew s >>= fun r' => cchk r >> k r') = cchk r >> (cnew s >>= k). Proof. @@ -315,7 +315,7 @@ do! rewrite bindE /=. by rewrite /bindS /= MS_mapE /= fmapE bindA. Qed. -Let cchkC T1 T2 (r1: loc T1) (r2: loc T2) : +Definition cchkC T1 T2 (r1: loc T1) (r2: loc T2) : cchk r1 >> cchk r2 = cchk r2 >> cchk r1. Proof. congr mkActo. @@ -342,7 +342,7 @@ do! rewrite bindE /=. by rewrite Hr1 Hc. Qed. -Let cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. +Definition cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. Proof. congr mkActo. apply/boolp.funext => st /=. @@ -352,7 +352,7 @@ case Hc: (coerce T v1) => [u1|] //. by rewrite bindE /= bindE /= Hr1 Hc. Qed. -Let cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type T2 -> M A) : +Definition cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type T2 -> M A) : cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). Proof. congr mkActo. @@ -381,7 +381,25 @@ do! rewrite bindE /=. by rewrite /bindS /= MS_mapE /= fmapE Hr1 Hc. Qed. -Let cchkget T (r : loc T) (A: UU0) (k : coq_type T -> M A) : +Definition cchknew T1 T2 (r1 : loc T1) (s : coq_type T2) (A : UU0) + (k1 k2 : loc T2 -> M A) : + (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cchk r1 >> (cnew s >>= k1) = cchk r1 >> (cnew s >>= k2). +Proof. +move=> Hk. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] //=. +case Hc: (coerce T1 v1) => [u1|] //. +rewrite bindE /= bindE /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= [in RHS]bindE /= [in RHS]bindE /= /bindS MS_mapE /=. +rewrite fmapE /= bindA /= 2!bindE /=. +by rewrite Hk // neq_ltn /= (nth_error_size Hr1). +Qed. + +Definition cchkget T (r : loc T) (A: UU0) (k : coq_type T -> M A) : cchk r >> (cget r >>= k) = cget r >>= k. Proof. congr mkActo. @@ -394,7 +412,7 @@ do! rewrite bindE /=. by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. Qed. -Let cgetchk T (r : loc T) (A: UU0) (k : coq_type T -> M A) : +Definition cgetchk T (r : loc T) (A: UU0) (k : coq_type T -> M A) : cget r >>= (fun s => cchk r >> k s) = cget r >>= k. Proof. congr mkActo. @@ -407,7 +425,7 @@ do! rewrite bindE /=. by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. Qed. -Let cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : +Definition cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : cchk r1 >> cput r2 s = cput r2 s >> cchk r1. Proof. congr mkActo. @@ -453,7 +471,7 @@ rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. by case: ml_type_eq_dec. Qed. -Let cchkput T (r : loc T) (s : coq_type T) : +Definition cchkput T (r : loc T) (s : coq_type T) : cchk r >> cput r s = cput r s. Proof. congr mkActo. @@ -471,7 +489,7 @@ case: ml_type_eq_dec => // HT. by rewrite -eq_rect_eq. Qed. -Let cputchk T (r : loc T) (s : coq_type T) : +Definition cputchk T (r : loc T) (s : coq_type T) : cput r s >> cchk r = cput r s. Proof. congr mkActo. @@ -487,7 +505,7 @@ rewrite nth_error_set_nth. by case: ml_type_eq_dec. Qed. -Let cputput T (r : loc T) (s s' : coq_type T) : +Definition cputput T (r : loc T) (s s' : coq_type T) : cput r s >> cput r s' = cput r s'. Proof. congr mkActo. @@ -504,7 +522,7 @@ rewrite -eq_rect_eq. by rewrite set_set_nth eqxx. Qed. -Let cputget T (r : loc T) (s : coq_type T) (A : UU0) (k : coq_type T -> M A) : +Definition cputget T (r : loc T) (s : coq_type T) (A : UU0) (k : coq_type T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. congr mkActo. @@ -524,7 +542,7 @@ case: ml_type_eq_dec => H /=; last by rewrite bindfailf. by rewrite -eq_rect_eq !bindretf. Qed. -Let cgetputchk T (r : loc T) : cget r >>= cput r = cchk r. +Definition cgetputchk T (r : loc T) : cget r >>= cput r = cchk r. Proof. congr mkActo. apply/boolp.funext => st /=. @@ -538,7 +556,7 @@ case: (ml_type_eq_dec T T) => H //. by rewrite -eq_rect_eq nth_error_set_nth_id. Qed. -Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : +Definition cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. congr mkActo. @@ -556,7 +574,7 @@ case: (ml_type_eq_dec T T) => H //. by rewrite -eq_rect_eq. Qed. -Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) +Definition cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) (k : coq_type T1 -> coq_type T2 -> M A) : cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). @@ -599,7 +617,7 @@ case: (ml_type_eq_dec T1 T1) => // H. by rewrite -eq_rect_eq. Qed. -Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) +Definition cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) (s2 : coq_type T2) (A : UU0) : loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. @@ -667,7 +685,7 @@ case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq set_set_nth (negbTE Hr). Qed. -Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) +Definition cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) (A : UU0) (k : coq_type T2 -> M A) : loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). @@ -708,7 +726,7 @@ case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq bindE. Qed. -Let crunret (A B : UU0) (m : M A) (s : B) : +Definition crunret (A B : UU0) (m : M A) (s : B) : crun m -> crun (m >> ret s) = Some s. Proof. case: m => T /= m. @@ -716,10 +734,10 @@ rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. by case Hm: (m [::]). Qed. -Let crunskip : crun skip = Some tt. +Definition crunskip : crun skip = Some tt. Proof. by []. Qed. -Let crunnew (A : UU0) T (m : M A) (s : coq_type T) : +Definition crunnew (A : UU0) T (m : M A) (s : coq_type T) : crun m -> crun (m >> cnew s). Proof. case: m => U /= m. @@ -746,6 +764,7 @@ exists cnew cget cput cchk crun. - exact cnewchk. - exact cchknewC. - exact cchkgetC. +- exact cchknew. - exact cchkget. - exact cgetchk. - exact cchkputC. From bc89acdc96dbb27f02b075ac9b33fb3c6602ffe6 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Wed, 1 Feb 2023 23:17:00 +0900 Subject: [PATCH 16/82] compress --- example_typed_store.v | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index c21c9960..011cb938 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -132,7 +132,7 @@ have : i <= n by []. elim: i x y => [|i IH] x y Hi. rewrite !subn0 => -[] -> ->. rewrite -/(fibo_rec n.+1). - under funext do under funext do rewrite [fibo_ref _ _ _]/= bindskipf. + under funext do under funext do rewrite /= bindskipf. under funext do rewrite cgetret. rewrite -cnewchk. under funext do rewrite cnewgetC. @@ -142,8 +142,7 @@ rewrite -(IH y (x + y) (ltnW Hi)); last first. subst x y. congr pair. case: n Hi {IH} => // n. - rewrite ltnS => Hi. - rewrite subSS. + rewrite subSS ltnS => Hi. by rewrite -addn2 -addn1 -addnBAC // -addnBAC // addn2 addn1. rewrite /=. under funext do under funext do rewrite !bindA. From 035a3941f67f815126e14c47eb12ff2520d1ad92 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Wed, 1 Feb 2023 23:38:49 +0900 Subject: [PATCH 17/82] add cchknew again... --- hierarchy.v | 4 ++++ monad_model.v | 31 +++++++++++++++++++++++++++---- 2 files changed, 31 insertions(+), 4 deletions(-) diff --git a/hierarchy.v b/hierarchy.v index b28ec49a..19a069b2 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1149,6 +1149,10 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cchkgetC : forall T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type M T2 -> M A), cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s) ; + cchknew : forall T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) + (k1 k2 : loc T2 -> M A), + (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cchk r1 >> (cnew s >>= k1) = cchk r1 >> (cnew s >>= k2) ; cchkget : forall T (r : loc T) (A: UU0) (k : coq_type M T -> M A), cchk r >> (cget r >>= k) = cget r >>= k ; cgetchk : forall T (r : loc T) (A: UU0) (k : coq_type M T -> M A), diff --git a/monad_model.v b/monad_model.v index 5b4828c7..c84bd216 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1695,9 +1695,13 @@ Lemma nth_error_set_nth_none m n a b : nth_error (set_nth def st n b) m = None. Proof. by elim: m st n => [|m IH] [|c st'] [|n] //=; apply IH. Qed. +Lemma nth_error_size n a : + nth_error st n = Some a -> n < size st. +Proof. by elim: n st => [|n IH] [|c st'] //= /IH. Qed. + Lemma nth_error_size_set_nth n a b : nth_error st n = Some a -> size (set_nth def st n b) = size st. -Proof. by elim: n st => [|n IH] [|c st'] //= /IH ->. Qed. +Proof. by rewrite size_set_nth => /nth_error_size /maxn_idPr. Qed. Lemma set_nth_rcons a b : set_nth def (rcons st a) (size st) b = rcons st b. @@ -1729,6 +1733,7 @@ Module MTypedStore := MonadTypedStore (MLtypes'). Import MLtypes'. Import MTypedStore. +Section typed_store. Definition acto (T : UU0) : UU0 := MS (seq (binding idfun)) [the monad of option_monad] T. Local Notation M := acto. @@ -1918,7 +1923,24 @@ case Hc: (coerce T v1) => [u1|] //. by rewrite bindE /= bindE /= Hr1 Hc. Qed. -Let cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type M T2 -> M A) : +Let cchknew T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) + (k1 k2 : loc T2 -> M A) : + (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cchk r1 >> (cnew s >>= k1) = cchk r1 >> (cnew s >>= k2). +Proof. +move=> Hk. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] //=. +case Hc: (coerce T1 v1) => [u1|] //. +rewrite bindE /= bindE /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= [in RHS]bindE /= [in RHS]bindE /= /bindS MS_mapE /=. +rewrite fmapE /= bindA /= 2!bindE /=. +by rewrite Hk // neq_ltn /= (nth_error_size Hr1). +Qed. + +Let cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type M T2 -> M A): cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). Proof. apply/boolp.funext => st /=. @@ -2281,16 +2303,17 @@ Qed. Canonical Structure isMonadTypedStoreModel := isMonadTypedStore.Build M cnewget cnewgetC cnewput cnewputC cputput cputget cgetputchk cgetget cgetC cputC cputgetC cnewchk cchknewC - cchkgetC cchkget cgetchk cchkputC cchkput cputchk cchkC cchkdup + cchkgetC cchknew cchkget cgetchk cchkputC cchkput cputchk cchkC cchkdup crunret crunskip crunnew. -(* Fails +(* Fails or diverges ? HB.instance Definition _ := isMonadTypedStore.Build M cnewget cnewgetC cnewput cnewputC cputput cputget cgetputchk cgetget cgetC cputC cputgetC cnewchk cchknewC cchkgetC cchkget cgetchk cchkputC cchkput cputchk cchkC cchkdup crunret crunskip crunnew. *) +End typed_store. End ModelTypedStore. (* TODO? From 4c2f68b57efc10f364bbb1f125a32de3b4b2f279 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 3 Feb 2023 13:19:34 +0900 Subject: [PATCH 18/82] small improvements --- example_typed_store.v | 4 ++-- monad_model.v | 2 +- typed_store_model.v | 16 ++++++++-------- 3 files changed, 11 insertions(+), 11 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 011cb938..299d77c1 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -81,7 +81,7 @@ Notation coq_type := (@MLtypes.coq_type M). Fixpoint fact_ref (r : loc ml_int) (n : nat) : M unit := if n is m.+1 then cget r >>= fun p => cput r (n * p) >> fact_ref r m - else Ret tt. + else skip. Theorem fact_ref_ok n : crun (cnew ml_int 1 >>= fun r => fact_ref r n >> cget r) = Some (fact_rec n). @@ -92,7 +92,7 @@ set s := 1. have smn : s * fact_rec m = fn by rewrite mul1n. elim: m s smn => [|m IH] s /= smn. rewrite /fact_ref -smn muln1. - under funext do rewrite bindretf. + under funext do rewrite bindskipf. by rewrite cgetret cnewget crunret // crunnew0. under funext do rewrite bindA. rewrite cnewget. diff --git a/monad_model.v b/monad_model.v index c84bd216..9f96125c 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1680,7 +1680,7 @@ Proof. by elim: n st => [|n IH] []. Qed. Lemma nth_error_set_nth_id n a : nth_error st n = Some a -> set_nth def st n a = st. -Proof. elim: n st =>[|n IH] [] //= b st'; by [case=> -> | move/IH ->]. Qed. +Proof. by elim: n st => [|n IH] [] //= b st'; [case=> -> | move/IH ->]. Qed. Lemma nth_error_set_nth_other m n a b : m != n -> diff --git a/typed_store_model.v b/typed_store_model.v index 6ea32762..b6334d54 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -90,6 +90,14 @@ Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. Definition associative : BindLaws.associative bind. Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. +(* Doesn't work +HB.instance Definition xyz := + isMonad_ret_bind.Build M left_neutral right_neutral associative. +*) +(* Since we couldn't build the instance, redefine notations *) +(*Local Notation "m >>= f" := (bind m f). +Local Notation "m >> f" := (bind m (fun=> f)).*) + Definition actm (A B : UU0) (f : A -> B) (m : M A) : M B := mkActo (actm _ _ f (ofActo m)). @@ -191,14 +199,6 @@ Canonical Structure MonadTS : Monad M := Monad.Class isMonadTS. Canonical Structure monadTS : monad := Monad.Pack isMonadTS. Set Universe Checking. -(* Doesn't work -HB.instance Definition xyz := - isMonad_ret_bind.Build M left_neutral right_neutral associative. -*) -(* Since we couldn't build the instance, redefine notations *) -(*Local Notation "m >>= f" := (bind m f). -Local Notation "m >> f" := (bind m (fun=> f)).*) - (* Make ml_type an eqType *) Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. Lemma ml_type_eqP : Equality.axiom ml_type_eqb. From d0eabbbe243d52d5ddb63124c2f56e12083f2e6d Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Wed, 8 Feb 2023 16:07:28 +0900 Subject: [PATCH 19/82] use eq_bind in example --- example_typed_store.v | 39 +++++++++++++++++++++------------------ 1 file changed, 21 insertions(+), 18 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 299d77c1..23a23dc2 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -65,15 +65,19 @@ Arguments cput {s} [T]. Arguments cchk {s} [T]. Arguments crun {s} [A]. -Lemma cgetret (M : typedStoreMonad) T : +Section lemmas. +Variable M : typedStoreMonad. + +Lemma cgetret T : @cget _ T = fun r => (cget r : M _) >>= Ret. Proof. by apply boolp.funext => r; rewrite bindmret. Qed. -Lemma crunnew0 (M : typedStoreMonad) T s : crun (cnew T s : M _). +Lemma crunnew0 T s : crun (cnew T s : M _). Proof. by rewrite -[cnew _ _]bindskipf crunnew // crunskip. Qed. -Lemma funext {T U : UU0} [f g : T -> U] : f =1 g -> (fun x => f x) = g. -Proof. exact: boolp.funext. Qed. +Lemma cnewgetret T s : cnew T s >>= @cget _ T = cnew T s >> Ret s :> M _. +Proof. by rewrite cgetret cnewget. Qed. +End lemmas. Section factorial. Variable M : typedStoreMonad. @@ -92,11 +96,11 @@ set s := 1. have smn : s * fact_rec m = fn by rewrite mul1n. elim: m s smn => [|m IH] s /= smn. rewrite /fact_ref -smn muln1. - under funext do rewrite bindskipf. - by rewrite cgetret cnewget crunret // crunnew0. -under funext do rewrite bindA. + under eq_bind do rewrite bindskipf. + by rewrite cnewgetret crunret // crunnew0. +under eq_bind do rewrite bindA. rewrite cnewget. -under funext do rewrite bindA. +under eq_bind do rewrite bindA. by rewrite cnewput IH // (mulnC m.+1) -mulnA. Qed. End factorial. @@ -132,10 +136,9 @@ have : i <= n by []. elim: i x y => [|i IH] x y Hi. rewrite !subn0 => -[] -> ->. rewrite -/(fibo_rec n.+1). - under funext do under funext do rewrite /= bindskipf. - under funext do rewrite cgetret. + under eq_bind do under eq_bind do rewrite /= bindskipf. rewrite -cnewchk. - under funext do rewrite cnewgetC. + under eq_bind do rewrite cgetret cnewgetC. by rewrite cnewget -bindA crunret // crunnew // crunnew0. rewrite subSS => -[] Hx Hy. rewrite -(IH y (x + y) (ltnW Hi)); last first. @@ -145,16 +148,16 @@ rewrite -(IH y (x + y) (ltnW Hi)); last first. rewrite subSS ltnS => Hi. by rewrite -addn2 -addn1 -addnBAC // -addnBAC // addn2 addn1. rewrite /=. -under funext do under funext do rewrite !bindA. +under eq_bind do under eq_bind do rewrite !bindA. rewrite -cnewchk. -under funext do rewrite cnewgetC. +under eq_bind do rewrite cnewgetC. rewrite cnewget. -under funext do under funext do rewrite !bindA. -under funext do rewrite cnewget. -under funext do under funext do rewrite !bindA. +under eq_bind do under eq_bind do rewrite !bindA. +under eq_bind do rewrite cnewget. +under eq_bind do under eq_bind do rewrite !bindA. rewrite -cnewchk. -under funext do rewrite cnewputC. +under eq_bind do rewrite cnewputC. rewrite cnewput. -by under funext do rewrite cnewput. +by under eq_bind do rewrite cnewput. Qed. End fibonacci. From 7942e1307ac97c4d0771a5714821aa3a7b926e54 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Wed, 8 Feb 2023 16:47:10 +0900 Subject: [PATCH 20/82] make loc abstract --- example_typed_store.v | 6 +++++- hierarchy.v | 10 +++------- monad_model.v | 14 ++++++++++++-- typed_store_model.v | 21 +++++++++++++++++++-- 4 files changed, 39 insertions(+), 12 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 23a23dc2..41bc8ef1 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -23,6 +23,8 @@ Inductive ml_type : Set := Inductive undef_t : Set := Undef. +Variant loc : ml_type -> Type := mkloc T : nat -> loc T. + Module MLtypes. Definition ml_type_eq_dec (T1 T2 : ml_type) : {T1=T2}+{T1<>T2}. revert T2; induction T1; destruct T2; @@ -37,7 +39,9 @@ Defined. Local Definition ml_type := ml_type. Local Definition undef := Undef. -Variant loc : ml_type -> Type := mkloc T : nat -> loc T. +Local Definition loc := loc. +Local Definition locT := [eqType of nat]. +Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. Section with_monad. Context [M : Type -> Type]. diff --git a/hierarchy.v b/hierarchy.v index 19a069b2..69cb19fa 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1080,19 +1080,15 @@ HB.structure Definition MonadPlusArray (S : UU0) (I : eqType) := Module Type MLTY. Parameter ml_type : Set. Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Variant loc : ml_type -> Type := mkloc T : nat -> loc T. +Parameter loc : ml_type -> Type. +Parameter locT : eqType. +Parameter loc_id : forall T, loc T -> locT. Parameter coq_type : forall M : Type -> Type, ml_type -> Type. End MLTY. Module MonadTypedStore (MLtypes : MLTY). Import MLtypes. -Record binding (M : Type -> Type) := - mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. -Arguments mkbind {M}. - -Definition loc_id {T} (r : loc T) : nat := let: mkloc _ n := r in n. - HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) of Monad M := { cnew : forall {T}, coq_type M T -> M (loc T) ; diff --git a/monad_model.v b/monad_model.v index 9f96125c..46a57b56 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1718,21 +1718,31 @@ Require Import JMeq. Module Type MLTYweak. Parameter ml_type : Set. Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Variant loc : ml_type -> Type := mkloc T : nat -> loc T. +Variant loc : ml_type -> Set := mkloc T : nat -> loc T. Parameter coq_type0 : ml_type -> Type. Parameter ml_undef : ml_type. Parameter undef : coq_type0 ml_undef. End MLTYweak. Module ModelTypedStore (MLtypes : MLTYweak). +Import MLtypes. Module MLtypes'. -Include MLtypes. +Definition ml_type := ml_type. +Definition ml_type_eq_dec := ml_type_eq_dec. Definition coq_type (M : UU0 -> UU0) := coq_type0. +Definition loc := loc. +Definition locT := [eqType of nat]. +Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. End MLtypes'. Module MTypedStore := MonadTypedStore (MLtypes'). Import MLtypes'. +Import MLtypes. Import MTypedStore. +Record binding (M : Type -> Type) := + mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. +Arguments mkbind {M}. + Section typed_store. Definition acto (T : UU0) : UU0 := MS (seq (binding idfun)) [the monad of option_monad] T. diff --git a/typed_store_model.v b/typed_store_model.v index b6334d54..f304e707 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -20,16 +20,33 @@ Unset Printing Implicit Defensive. Local Open Scope monae_scope. Module Type MLTYdef. -Include MLTY. +Parameter ml_type : Set. +Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. +Variant loc : ml_type -> Type := mkloc T : nat -> loc T. +Parameter coq_type : forall M : Type -> Type, ml_type -> Type. Parameter ml_undef : ml_type. Parameter undef : forall M, coq_type M ml_undef. End MLTYdef. Module ModelTypedStore (MLtypes : MLTYdef). -Module MTypedStore := MonadTypedStore (MLtypes). +Import MLtypes. +Module MLtypes'. +Definition ml_type := ml_type. +Definition ml_type_eq_dec := ml_type_eq_dec. +Definition coq_type := coq_type. +Definition loc := loc. +Definition locT := [eqType of nat]. +Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. +End MLtypes'. +Module MTypedStore := MonadTypedStore (MLtypes'). +Import MLtypes'. Import MLtypes. Import MTypedStore. +Record binding (M : Type -> Type) := + mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. +Arguments mkbind {M}. + #[bypass_check(positivity)] Inductive acto : UU0 -> UU0 := mkActo : forall T : UU0, From 65d9c2cffc5bfb6f0efa74ef4048a9d47353af1b Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Thu, 9 Feb 2023 19:16:07 +0900 Subject: [PATCH 21/82] make cchk a derived operation --- example_typed_store.v | 29 +-- hierarchy.v | 157 ++++++++++---- monad_model.v | 444 +++++++++++++------------------------- typed_store_model.v | 491 ++++++++++++++---------------------------- 4 files changed, 435 insertions(+), 686 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 41bc8ef1..52c28312 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -46,7 +46,7 @@ Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. Section with_monad. Context [M : Type -> Type]. -Local (* Generated type translation function *) +(* Generated type translation function *) Fixpoint coq_type (T : ml_type) : Type := match T with | ml_int => nat @@ -64,25 +64,6 @@ Module IMonadTS := MonadTypedStore (MLtypes). Import MLtypes. Import IMonadTS. -Arguments cget {s} [T]. -Arguments cput {s} [T]. -Arguments cchk {s} [T]. -Arguments crun {s} [A]. - -Section lemmas. -Variable M : typedStoreMonad. - -Lemma cgetret T : - @cget _ T = fun r => (cget r : M _) >>= Ret. -Proof. by apply boolp.funext => r; rewrite bindmret. Qed. - -Lemma crunnew0 T s : crun (cnew T s : M _). -Proof. by rewrite -[cnew _ _]bindskipf crunnew // crunskip. Qed. - -Lemma cnewgetret T s : cnew T s >>= @cget _ T = cnew T s >> Ret s :> M _. -Proof. by rewrite cgetret cnewget. Qed. -End lemmas. - Section factorial. Variable M : typedStoreMonad. Notation coq_type := (@MLtypes.coq_type M). @@ -142,7 +123,7 @@ elim: i x y => [|i IH] x y Hi. rewrite -/(fibo_rec n.+1). under eq_bind do under eq_bind do rewrite /= bindskipf. rewrite -cnewchk. - under eq_bind do rewrite cgetret cnewgetC. + under eq_bind do rewrite -cgetret cchknewget. by rewrite cnewget -bindA crunret // crunnew // crunnew0. rewrite subSS => -[] Hx Hy. rewrite -(IH y (x + y) (ltnW Hi)); last first. @@ -154,13 +135,13 @@ rewrite -(IH y (x + y) (ltnW Hi)); last first. rewrite /=. under eq_bind do under eq_bind do rewrite !bindA. rewrite -cnewchk. -under eq_bind do rewrite cnewgetC. +under eq_bind do rewrite cchknewget. rewrite cnewget. under eq_bind do under eq_bind do rewrite !bindA. under eq_bind do rewrite cnewget. under eq_bind do under eq_bind do rewrite !bindA. -rewrite -cnewchk. -under eq_bind do rewrite cnewputC. +rewrite -[in LHS]cnewchk. +under eq_bind do rewrite cchknewput. rewrite cnewput. by under eq_bind do rewrite cnewput. Qed. diff --git a/hierarchy.v b/hierarchy.v index 69cb19fa..0371db39 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1094,36 +1094,38 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) cnew : forall {T}, coq_type M T -> M (loc T) ; cget : forall {T}, loc T -> M (coq_type M T) ; cput : forall {T}, loc T -> coq_type M T -> M unit ; - cchk : forall {T}, loc T -> M unit ; crun : forall {A : UU0}, M A -> option A ; (* execute in empty store *) cnewget : forall T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A), cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s) ; - cnewgetC : - forall T T' (r : loc T) (s : coq_type M T') A - (k : loc T' -> coq_type M T -> M A), - cchk r >> (cnew s >>= (fun r' => cget r >>= k r')) = - cget r >>= (fun u => cnew s >>= (fun r' => k r' u)) ; cnewput : forall T (s t : coq_type M T) A (k : loc T -> M A), cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k ; - cnewputC : - forall T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A - (k : loc T' -> M A), - cchk r >> (cnew s' >>= fun r' => cput r s >> k r') = - cput r s >> (cnew s' >>= k) ; - cputput : forall T (r : loc T) (s s' : coq_type M T), - cput r s >> cput r s' = cput r s' ; - cputget : - forall T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A), - cput r s >> (cget r >>= k) = cput r s >> k s ; - cgetputchk : forall T (r : loc T), cget r >>= cput r = cchk r ; + cgetput : forall T (r : loc T) (s : coq_type M T), + cget r >> cput r s = cput r s ; + cgetputskip : forall T (r : loc T), cget r >>= cput r = cget r >> skip ; cgetget : forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s ; + cputget : + forall T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A), + cput r s >> (cget r >>= k) = cput r s >> k s ; + cputput : forall T (r : loc T) (s s' : coq_type M T), + cput r s >> cput r s' = cput r s' ; cgetC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) (k : coq_type M T1 -> coq_type M T2 -> M A), cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)) ; + cgetnewD : + forall T T' (r : loc T) (s : coq_type M T') A + (k : loc T' -> coq_type M T -> coq_type M T -> M A), + cget r >>= (fun u => cnew s >>= (fun r' => cget r >>= k r' u)) = + cget r >>= (fun u => cnew s >>= (fun r' => k r' u u)) ; + cgetnewE : forall T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) + (k1 k2 : loc T2 -> M A), + (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2) ; + cgetputC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2), + cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip ; cputC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) (s2 : coq_type M T2) (A : UU0), @@ -1135,34 +1137,11 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v) ; - cnewchk : - forall T (s : coq_type M T) (A : UU0) (k : loc T -> M A), - cnew s >>= (fun r => cchk r >> k r) = cnew s >>= k ; - cchknewC : - forall T1 T2 (r : loc T1) (s : coq_type M T2) (A : UU0) (k : loc T2 -> M A), - cchk r >> (cnew s >>= fun r' => cchk r >> k r') = - cchk r >> (cnew s >>= k) ; - cchkgetC : - forall T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type M T2 -> M A), - cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s) ; - cchknew : forall T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) - (k1 k2 : loc T2 -> M A), - (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> - cchk r1 >> (cnew s >>= k1) = cchk r1 >> (cnew s >>= k2) ; - cchkget : forall T (r : loc T) (A: UU0) (k : coq_type M T -> M A), - cchk r >> (cget r >>= k) = cget r >>= k ; - cgetchk : forall T (r : loc T) (A: UU0) (k : coq_type M T -> M A), - cget r >>= (fun s => cchk r >> k s) = cget r >>= k ; - cchkputC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2), - cchk r1 >> cput r2 s = cput r2 s >> cchk r1 ; - cchkput : forall T (r : loc T) (s : coq_type M T), - cchk r >> cput r s = cput r s ; - cputchk : forall T (r : loc T) (s : coq_type M T), - cput r s >> cchk r = cput r s ; - cchkC : forall T1 T2 (r1: loc T1) (r2: loc T2), - cchk r1 >> cchk r2 = cchk r2 >> cchk r1 ; - cchkdup : forall T (r : loc T), - cchk r >> cchk r = cchk r ; + cputnewC : + forall T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A + (k : loc T' -> M A), + cget r >> (cnew s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew s' >>= k) ; crunret : forall (A B : UU0) (m : M A) (s : B), crun m -> crun (m >> Ret s) = Some s ; crunskip : @@ -1174,6 +1153,94 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) #[short(type=typedStoreMonad)] HB.structure Definition MonadTypedStore := { M of isMonadTypedStore M & isMonad M & isFunctor M }. + +Arguments cnew {s}. +Arguments cget {s} [T]. +Arguments cput {s} [T]. +Arguments crun {s} [A]. + +Section cchk. +Variable M : typedStoreMonad. +Definition cchk T (r : loc T) : M unit := cget r >> skip. + +Lemma cnewchk T s (A : UU0) (k : loc T -> M A) : + cnew T s >>= (fun r => cchk r >> k r) = cnew T s >>= k. +Proof. +under eq_bind do rewrite bindA. +rewrite cnewget. +by under eq_bind do rewrite bindskipf. +Qed. + +Lemma cchknewC T1 T2 (r : loc T1) s (A : UU0) (k : loc T2 -> M A) : + cchk r >> (cnew T2 s >>= fun r' => cchk r >> k r') = + cchk r >> (cnew T2 s >>= k). +Proof. +rewrite !(bindA,bindskipf). +under eq_bind do under eq_bind do rewrite bindA. +rewrite cgetnewD. +by under eq_bind do under eq_bind do rewrite bindskipf. +Qed. + +Lemma cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) + (k: coq_type M T2 -> M A) : + cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). +Proof. +under [in RHS]eq_bind do rewrite bindA bindskipf. +by rewrite !(bindA,bindskipf) cgetC. +Qed. + +Lemma cchknewE T1 T2 (r1 : loc T1) s (A : UU0) (k1 k2 : loc T2 -> M A) : + (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cchk r1 >> (cnew T2 s >>= k1) = cchk r1 >> (cnew T2 s >>= k2). +Proof. move=> Hk; rewrite !(bindA,bindskipf); exact: cgetnewE. Qed. + +Lemma cchknewget T T' (r : loc T) s (A : UU0) k : + cchk r >> (cnew T' s >>= fun r' => cget r >>= k r') = + cget r >>= (fun u => cnew T' s >>= k ^~ u) :> M A. +Proof. by rewrite bindA bindskipf cgetnewD. Qed. + +Lemma cchknewput T T' (r : loc T) s s' (A : UU0) k : + cchk r >> (cnew T' s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew T' s' >>= k) :> M A. +Proof. by rewrite bindA bindskipf cputnewC. Qed. + +Lemma cchkget T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : + cchk r >> (cget r >>= k) = cget r >>= k. +Proof. by rewrite bindA bindskipf cgetget. Qed. + +Lemma cgetchk T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : + cget r >>= (fun s => cchk r >> k s) = cget r >>= k. +Proof. under eq_bind do rewrite bindA bindskipf; by rewrite cgetget. Qed. + +Lemma cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : + cchk r1 >> cput r2 s = cput r2 s >> cchk r1. +Proof. by rewrite bindA bindskipf cgetputC bindA. Qed. + +Lemma cchkput T (r : loc T) (s : coq_type M T) : + cchk r >> cput r s = cput r s. +Proof. by rewrite bindA bindskipf cgetput. Qed. + +Lemma cputchk T (r : loc T) (s : coq_type M T) : + cput r s >> cchk r = cput r s. +Proof. by rewrite cputget bindmskip. Qed. + +Lemma cchkC T1 T2 (r1: loc T1) (r2: loc T2) : + cchk r1 >> cchk r2 = cchk r2 >> cchk r1. +Proof. by rewrite !(bindA,bindskipf) cgetC. Qed. + +Lemma cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. +Proof. by rewrite bindA bindskipf cgetget. Qed. + +Lemma cgetret T (r : loc T) : cget r >>= Ret = cget r :> M _. +Proof. by rewrite bindmret. Qed. + +Lemma crunnew0 T s : crun (cnew T s : M _). +Proof. by rewrite -(bindskipf (cnew T s)) crunnew // crunskip. Qed. + +Lemma cnewgetret T s : cnew T s >>= @cget _ _ = cnew T s >> Ret s :> M _. +Proof. under eq_bind do rewrite -cgetret; by rewrite cnewget. Qed. +End cchk. +Arguments cchk {M} [T]. End MonadTypedStore. HB.mixin Record isMonadTrace (T : UU0) (M : UU0 -> UU0) of Monad M := diff --git a/monad_model.v b/monad_model.v index 46a57b56..29921cd5 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1778,12 +1778,6 @@ Definition cput T (r : loc T) (v : coq_type M T) : M unit := else fail else fail. -Definition cchk T (r : loc T) : M unit := - fun st => - if nth_error st (loc_id r) is Some (mkbind T' u) then - if coerce T u is Some _ then Ret (tt, st) else fail - else fail. - Definition crun (A : UU0) (m : M A) : option A := match m nil with | inl _ => None @@ -1813,26 +1807,6 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. Qed. -Let cnewgetC T T' (r : loc T) (s : coq_type M T') A - (k : loc T' -> coq_type M T -> M A) : - cchk r >> (cnew s >>= (fun r' => cget r >>= k r')) = - cget r >>= (fun u => cnew s >>= fun r' => k r' u). -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T1. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite (nth_error_rcons_some _ Hr). -rewrite /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq. -Qed. - Let cnewput T (s t : coq_type M T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. @@ -1847,215 +1821,16 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. Qed. -Let cnewputC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A - (k : loc T' -> M A) : - cchk r >> (cnew s' >>= fun r' => cput r s >> k r') = - cput r s >> (cnew s' >>= k). -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cput. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. -subst T1. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite (nth_error_rcons_some _ Hr). -rewrite /coerce. -case: ml_type_eq_dec => // H. -rewrite -eq_rect_eq. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /=. -rewrite (nth_error_size_set_nth _ _ Hr). -by rewrite (nth_error_set_nth_rcons _ _ _ Hr). -Qed. - -Let cnewchk T (s : coq_type M T) (A : UU0) (k : loc T -> M A) : - cnew s >>= (fun r => cchk r >> k r) = cnew s >>= k. -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite /cchk /cnew nth_error_rcons_size /coerce. -by case: ml_type_eq_dec. -Qed. - -Let cchknewC T1 T2 (r : loc T1) (s : coq_type M T2) (A : UU0) - (k : loc T2 -> M A) : - cchk r >> (cnew s >>= fun r' => cchk r >> k r') = cchk r >> (cnew s >>= k). +Let cgetput T (r : loc T) (s : coq_type M T) : + cget r >> cput r s = cput r s. Proof. apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cnew. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -case Hc: (coerce T1 u) => [u'|] //. -rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /= bindA. -rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /= bindA. -rewrite (nth_error_rcons_some _ Hr) Hc. -do! rewrite bindE /=. -by rewrite /bindS /= MS_mapE /= fmapE bindA. -Qed. - -Let cchkC T1 T2 (r1: loc T1) (r2: loc T2) : - cchk r1 >> cchk r2 = cchk r2 >> cchk r1. -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. - rewrite bindfailf. - case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; - last by rewrite bindfailf. - case Hc: (coerce T2 v2) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= Hr1. -case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. - rewrite bindfailf. - case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= Hr2. -case Hc: (coerce T1 v1) => [u1|]; last first. - rewrite bindfailf. - case Hc2: (coerce T2 v2) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= Hr1 Hc. -rewrite bindE /= bindE /= Hr2. -case Hc2: (coerce T2 v2) => [u2|]; last by rewrite !bindfailf. -do! rewrite bindE /=. -by rewrite Hr1 Hc. -Qed. - -Let cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk. -case Hr1 : (nth_error st (loc_id r)) => [[T1' v1]|] //=. -case Hc: (coerce T v1) => [u1|] //. -by rewrite bindE /= bindE /= Hr1 Hc. -Qed. - -Let cchknew T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) - (k1 k2 : loc T2 -> M A) : - (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> - cchk r1 >> (cnew s >>= k1) = cchk r1 >> (cnew s >>= k2). -Proof. -move=> Hk. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] //=. -case Hc: (coerce T1 v1) => [u1|] //. -rewrite bindE /= bindE /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= [in RHS]bindE /= [in RHS]bindE /= /bindS MS_mapE /=. -rewrite fmapE /= bindA /= 2!bindE /=. -by rewrite Hk // neq_ltn /= (nth_error_size Hr1). -Qed. - -Let cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type M T2 -> M A): - cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. - rewrite bindfailf. - case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; - last by rewrite bindfailf. - case Hc: (coerce T2 v2) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1 bindfailf. -case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. - rewrite bindfailf. - case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr2 bindfailf. -case Hc: (coerce T1 v1) => [u1|]; last first. - rewrite bindfailf. - case Hc2: (coerce T2 v2) => [u|]; last by rewrite bindfailf. - rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1 Hc. - by rewrite bindfailf. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite Hr2. -case Hc2: (coerce T2 v2) => [u2|]; last by rewrite !bindfailf. -do! rewrite bindE /=. -by rewrite /bindS /= MS_mapE /= fmapE Hr1 Hc. -Qed. - -Let cchkget T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : - cchk r >> (cget r >>= k) = cget r >>= k. -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -case Hc: (coerce T u) => [u'|] //. -do! rewrite bindE /=. -by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. -Qed. - -Let cgetchk T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : - cget r >>= (fun s => cchk r >> k s) = cget r >>= k. -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cchk. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -case Hc: (coerce T u) => [u'|] //. -do! rewrite bindE /=. -by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. -Qed. - -Let cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : - cchk r1 >> cput r2 s = cput r2 s >> cchk r1. -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cput. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. - rewrite bindfailf. - case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; - last by rewrite bindfailf. - case Hc: (coerce T2' s) => [u|]; last by rewrite bindfailf. - rewrite bindE /= /bindS /= bindE /=. - by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). -case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. - rewrite bindfailf. - case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= Hr2. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec T1' T1) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec T2 T2') => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE/= bindE/=. - case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. - rewrite -Hr nth_error_set_nth /coerce. - move: Hr2; rewrite -Hr Hr1 => -[] HT2. - subst T1'. - by case: ml_type_eq_dec. - rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. - by case: ml_type_eq_dec. -rewrite bindE /= bindE /= Hr2 {1}/coerce. -case: ml_type_eq_dec => HT2 //. -subst T1' T2'. -rewrite -!eq_rect_eq. -rewrite bindE /= bindE /=. -case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. - rewrite -Hr nth_error_set_nth /coerce. - move: Hr2; rewrite -Hr Hr1 => -[] HT2. - subst T2. - by case: ml_type_eq_dec. -rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. -by case: ml_type_eq_dec. -Qed. - -Let cchkput T (r : loc T) (s : coq_type M T) : - cchk r >> cput r s = cput r s. -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cput. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. rewrite /coerce. case: ml_type_eq_dec => HT; last by case: ml_type_eq_dec => // /esym //. subst T'. +rewrite -eq_rect_eq. case: ml_type_eq_dec => // HT. rewrite -eq_rect_eq {HT}. do! rewrite bindE /=. @@ -2064,38 +1839,38 @@ case: ml_type_eq_dec => // HT. by rewrite -eq_rect_eq. Qed. -Let cputchk T (r : loc T) (s : coq_type M T) : - cput r s >> cchk r = cput r s. +Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. Proof. apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cchk. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. rewrite /coerce. -case: ml_type_eq_dec => HT //. -subst T'. -rewrite -!eq_rect_eq. -do! rewrite bindE /=. -rewrite nth_error_set_nth. -by case: ml_type_eq_dec. +case: (ml_type_eq_dec T1 T) => H //. +subst T1. +rewrite bindE /= bindE /= Hr. +case: (ml_type_eq_dec T T) => H //. +by rewrite -eq_rect_eq nth_error_set_nth_id. Qed. -Let cputput T (r : loc T) (s s' : coq_type M T) : - cput r s >> cput r s' = cput r s'. +Let cgetget T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A) : + cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. -apply/boolp.funext => st. -case: r s s' => {}T n s s' /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. -case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T'. -rewrite !bindretf /= nth_error_set_nth /coerce. -case: ml_type_eq_dec => // H. -rewrite -eq_rect_eq. -by rewrite set_set_nth eqxx. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite /coerce. +case: (ml_type_eq_dec T1 T) => H //. +subst T1. +rewrite bindE /= bindE /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite Hr. +case: (ml_type_eq_dec T T) => H //. +by rewrite -eq_rect_eq. Qed. -Let cputget T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A) : +Let cputget T (r: loc T) (s: coq_type M T) (A: UU0) (k: coq_type M T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. apply/boolp.funext => st /=. @@ -2114,34 +1889,20 @@ case: ml_type_eq_dec => H /=; last by rewrite bindfailf. by rewrite -eq_rect_eq !bindretf. Qed. -Let cgetputchk T (r : loc T) : cget r >>= cput r = cchk r. -Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput /cchk. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite /coerce. -case: (ml_type_eq_dec T1 T) => H //. -subst T1. -rewrite bindE /= bindE /= Hr. -case: (ml_type_eq_dec T T) => H //. -by rewrite -eq_rect_eq nth_error_set_nth_id. -Qed. - -Let cgetget T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A) : - cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. +Let cputput T (r : loc T) (s s' : coq_type M T) : + cput r s >> cput r s' = cput r s'. Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite /coerce. -case: (ml_type_eq_dec T1 T) => H //. -subst T1. -rewrite bindE /= bindE /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite Hr. -case: (ml_type_eq_dec T T) => H //. -by rewrite -eq_rect_eq. +apply/boolp.funext => st. +case: r s s' => {}T n s s' /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. +case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T'. +rewrite !bindretf /= nth_error_set_nth /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +by rewrite set_set_nth eqxx. Qed. Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) @@ -2186,6 +1947,88 @@ case: (ml_type_eq_dec T1 T1) => // H. by rewrite -eq_rect_eq. Qed. +Let cgetnewD T T' (r : loc T) (s : coq_type M T') A + (k : loc T' -> coq_type M T -> coq_type M T -> M A) : + cget r >>= (fun u => cnew s >>= fun r' => cget r >>= k r' u) = + cget r >>= (fun u => cnew s >>= fun r' => k r' u u). +Proof. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T1. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite (nth_error_rcons_some _ Hr). +rewrite /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cgetnewE T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) + (k1 k2 : loc T2 -> M A) : + (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2). +Proof. +move=> Hk. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] //=. +case Hc: (coerce T1 v1) => [u1|] //. +rewrite bindE /= bindE /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= [in RHS]bindE /= [in RHS]bindE /= /bindS MS_mapE /=. +rewrite fmapE /= bindA /= 2!bindE /=. +by rewrite Hk // neq_ltn /= (nth_error_size Hr1). +Qed. + +Let cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : + cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip. +Proof. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= 2!bindA /= /cget /cput. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. + rewrite bindfailf. + case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; + last by rewrite bindfailf. + case Hc: (coerce T2' s) => [u|]; last by rewrite bindfailf. + rewrite bindE /= /bindS /= bindE /=. + by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). +case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. + rewrite bindfailf. + case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= Hr2. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec T1' T1) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec T2 T2') => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE/= bindE/=. + case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. + rewrite -Hr nth_error_set_nth /coerce. + move: Hr2; rewrite -Hr Hr1 => -[] HT2. + subst T1'. + by case: ml_type_eq_dec. + rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. + by case: ml_type_eq_dec. +rewrite bindE /= bindE /= Hr2 {1}/coerce. +case: ml_type_eq_dec => HT2 //. +subst T1' T2'. +rewrite -!eq_rect_eq. +rewrite bindE /= bindE /=. +case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. + rewrite -Hr nth_error_set_nth /coerce. + move: Hr2; rewrite -Hr Hr1 => -[] HT2. + subst T2. + by case: ml_type_eq_dec. +rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. +by case: ml_type_eq_dec. +Qed. + Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) (s2 : coq_type M T2) (A : UU0) : loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> @@ -2293,6 +2136,30 @@ case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq bindE. Qed. +Let cputnewC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A + (k : loc T' -> M A) : + cget r >> (cnew s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew s' >>= k). +Proof. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. +subst T1. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite (nth_error_rcons_some _ Hr). +rewrite /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /=. +rewrite (nth_error_size_set_nth _ _ Hr). +by rewrite (nth_error_set_nth_rcons _ _ _ Hr). +Qed. + Let crunret (A B : UU0) (m : M A) (s : B) : crun m -> crun (m >> Ret s) = Some s. Proof. @@ -2311,17 +2178,16 @@ by case Hm: (m [::]). Qed. Canonical Structure isMonadTypedStoreModel := - isMonadTypedStore.Build M cnewget cnewgetC cnewput cnewputC cputput - cputget cgetputchk cgetget cgetC cputC cputgetC cnewchk cchknewC - cchkgetC cchknew cchkget cgetchk cchkputC cchkput cputchk cchkC cchkdup - crunret crunskip crunnew. + isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip + cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC + cputgetC cputnewC crunret crunskip crunnew. -(* Fails or diverges ? +(* The elpi tactic/command HB.instance failed without giving a specific + error message. HB.instance Definition _ := - isMonadTypedStore.Build M cnewget cnewgetC cnewput cnewputC cputput - cputget cgetputchk cgetget cgetC cputC cputgetC cnewchk cchknewC - cchkgetC cchkget cgetchk cchkputC cchkput cputchk cchkC cchkdup - crunret crunskip crunnew. + isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip + cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC + cputgetC cputnewC crunret crunskip crunnew. *) End typed_store. End ModelTypedStore. diff --git a/typed_store_model.v b/typed_store_model.v index f304e707..153d60dc 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -235,34 +235,12 @@ apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite nth_error_rcons_size. -rewrite /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq. -Qed. - -Definition cnewgetC T T' (r : loc T) (s : coq_type T') A - (k : loc T' -> coq_type T -> M A) : - cchk r >> (cnew s >>= (fun r' => cget r >>= k r')) = - cget r >>= (fun u => cnew s >>= fun r' => k r' u). -Proof. -congr mkActo. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T1. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite (nth_error_rcons_some _ Hr). -rewrite /coerce. +rewrite nth_error_rcons_size /coerce. case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. Qed. -Definition cnewput T (s t : coq_type T) A (k : loc T -> M A) : +Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. congr mkActo. @@ -270,144 +248,176 @@ apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE bindA. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= [in RHS]bindE /=. +rewrite [in RHS]bindE /= [in RHS]bindE /= /cput. rewrite nth_error_rcons_size. rewrite /coerce. case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. Qed. -Definition cnewputC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A - (k : loc T' -> M A) : - cchk r >> (cnew s' >>= fun r' => cput r s >> k r') = - cput r s >> (cnew s' >>= k). +Let cgetput T (r : loc T) (s : coq_type T) : + cget r >> cput r s = cput r s. Proof. congr mkActo. apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. -subst T1. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite (nth_error_rcons_some _ Hr). +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. rewrite /coerce. -case: ml_type_eq_dec => // H. +case: ml_type_eq_dec => HT; last by case: ml_type_eq_dec => // /esym //. +subst T'. rewrite -eq_rect_eq. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /=. -rewrite (nth_error_size_set_nth _ _ Hr). -by rewrite (nth_error_set_nth_rcons _ _ _ Hr). +case: ml_type_eq_dec => // HT. +rewrite -eq_rect_eq {HT}. +do! rewrite bindE /=. +rewrite Hr. +case: ml_type_eq_dec => // HT. +by rewrite -eq_rect_eq. Qed. -Definition cnewchk T (s : coq_type T) (A : UU0) (k : loc T -> M A) : - cnew s >>= (fun r => cchk r >> k r) = cnew s >>= k. +Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. Proof. congr mkActo. apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite nth_error_rcons_size. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. rewrite /coerce. -by case: ml_type_eq_dec. +case: (ml_type_eq_dec T1 T) => H //. +subst T1. +rewrite bindE /= bindE /= Hr. +case: (ml_type_eq_dec T T) => H //. +by rewrite -eq_rect_eq nth_error_set_nth_id. Qed. -Definition cchknewC T1 T2 (r : loc T1) (s : coq_type T2) (A : UU0) - (k : loc T2 -> M A) : - cchk r >> (cnew s >>= fun r' => cchk r >> k r') = cchk r >> (cnew s >>= k). +Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : + cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -case Hc: (coerce T1 u) => [u'|] //. -rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /= bindA. -rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /= bindA. -rewrite (nth_error_rcons_some _ Hr) Hc. -do! rewrite bindE /=. -by rewrite /bindS /= MS_mapE /= fmapE bindA. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite /coerce. +case: (ml_type_eq_dec T1 T) => H //. +subst T1. +rewrite bindE /= bindE /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite Hr. +case: (ml_type_eq_dec T T) => H //. +by rewrite -eq_rect_eq. Qed. -Definition cchkC T1 T2 (r1: loc T1) (r2: loc T2) : - cchk r1 >> cchk r2 = cchk r2 >> cchk r1. +Let cputget T (r: loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : + cput r s >> (cget r >>= k) = cput r s >> k s. Proof. congr mkActo. apply/boolp.funext => st /=. +case: r s k => {}T n s k /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. - rewrite bindfailf. - case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; - last by rewrite bindfailf. - case Hc: (coerce T2 v2) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= Hr1. -case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. - rewrite bindfailf. - case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= Hr2. -case Hc: (coerce T1 v1) => [u1|]; last first. - rewrite bindfailf. - case Hc2: (coerce T2 v2) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= Hr1 Hc. -rewrite bindE /= bindE /= Hr2. -case Hc2: (coerce T2 v2) => [u2|]; last by rewrite !bindfailf. -do! rewrite bindE /=. -by rewrite Hr1 Hc. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. +case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T'. +rewrite !bindretf /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite nth_error_set_nth. +rewrite /coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +by rewrite -eq_rect_eq !bindretf. Qed. -Definition cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. +Let cputput T (r : loc T) (s s' : coq_type T) : + cput r s >> cput r s' = cput r s'. Proof. congr mkActo. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr1 : (nth_error st (loc_id r)) => [[T1' v1]|] //=. -case Hc: (coerce T v1) => [u1|] //. -by rewrite bindE /= bindE /= Hr1 Hc. +apply/boolp.funext => st. +case: r s s' => {}T n s s' /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. +case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T'. +rewrite !bindretf /= nth_error_set_nth /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +by rewrite set_set_nth eqxx. Qed. -Definition cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) (k: coq_type T2 -> M A) : - cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). +Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) + (k : coq_type T1 -> coq_type T2 -> M A) : + cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = + cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). Proof. congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. +case Hr1: (nth_error _ _) => [[T1' u]|]; last first. rewrite bindfailf. - case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; - last by rewrite bindfailf. - case Hc: (coerce T2 v2) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1 bindfailf. -case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. + case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. + case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. + rewrite !bindE /=. + by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. +case Hr2: (nth_error _ _) => [[T2' v]|]; last first. rewrite bindfailf. - case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr2 bindfailf. -case Hc: (coerce T1 v1) => [u1|]; last first. + case Hc: (coerce T1 u) => [v|]; last by rewrite bindfailf. + rewrite !bindE /=. + by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr2 bindfailf. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec _ _) => HT1; last first. rewrite bindfailf. - case Hc2: (coerce T2 v2) => [u|]; last by rewrite bindfailf. - rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1 Hc. - by rewrite bindfailf. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite Hr2. -case Hc2: (coerce T2 v2) => [u2|]; last by rewrite !bindfailf. -do! rewrite bindE /=. -by rewrite /bindS /= MS_mapE /= fmapE Hr1 Hc. + case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE /=. + rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr1 {1}/coerce. + case: (ml_type_eq_dec _ _) => //; by rewrite bindfailf. +subst T1'. +rewrite -eq_rect_eq. +rewrite bindE /=. +rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr2 {1}/coerce. +case: (ml_type_eq_dec T2' T2) => HT2; last by rewrite !bindfailf. +subst T2'. +rewrite -eq_rect_eq. +rewrite !bindE /=. +rewrite !bindE /=. +rewrite /bindS MS_mapE /= fmapE /= bindA /= Hr1 /coerce. +case: (ml_type_eq_dec T1 T1) => // H. +by rewrite -eq_rect_eq. Qed. -Definition cchknew T1 T2 (r1 : loc T1) (s : coq_type T2) (A : UU0) +Let cgetnewD T T' (r : loc T) (s : coq_type T') A + (k : loc T' -> coq_type T -> coq_type T -> M A) : + cget r >>= (fun u => cnew s >>= fun r' => cget r >>= k r' u) = + cget r >>= (fun u => cnew s >>= fun r' => k r' u u). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T1. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite (nth_error_rcons_some _ Hr). +rewrite /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cgetnewE T1 T2 (r1 : loc T1) (s : coq_type T2) (A : UU0) (k1 k2 : loc T2 -> M A) : (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> - cchk r1 >> (cnew s >>= k1) = cchk r1 >> (cnew s >>= k2). + cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2). Proof. move=> Hk. congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] //=. case Hc: (coerce T1 v1) => [u1|] //. rewrite bindE /= bindE /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. @@ -416,39 +426,13 @@ rewrite fmapE /= bindA /= 2!bindE /=. by rewrite Hk // neq_ltn /= (nth_error_size Hr1). Qed. -Definition cchkget T (r : loc T) (A: UU0) (k : coq_type T -> M A) : - cchk r >> (cget r >>= k) = cget r >>= k. +Let cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : + cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip. Proof. congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -case Hc: (coerce T u) => [u'|] //. -do! rewrite bindE /=. -by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. -Qed. - -Definition cgetchk T (r : loc T) (A: UU0) (k : coq_type T -> M A) : - cget r >>= (fun s => cchk r >> k s) = cget r >>= k. -Proof. -congr mkActo. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -case Hc: (coerce T u) => [u'|] //. -do! rewrite bindE /=. -by rewrite /bindS /= MS_mapE /= fmapE Hr Hc. -Qed. - -Definition cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : - cchk r1 >> cput r2 s = cput r2 s >> cchk r1. -Proof. -congr mkActo. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= 2!bindA /= /cget /cput. case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. rewrite bindfailf. case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; @@ -488,153 +472,7 @@ rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. by case: ml_type_eq_dec. Qed. -Definition cchkput T (r : loc T) (s : coq_type T) : - cchk r >> cput r s = cput r s. -Proof. -congr mkActo. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -rewrite /coerce. -case: ml_type_eq_dec => HT; last by case: ml_type_eq_dec => // /esym //. -subst T'. -case: ml_type_eq_dec => // HT. -rewrite -eq_rect_eq {HT}. -do! rewrite bindE /=. -rewrite Hr. -case: ml_type_eq_dec => // HT. -by rewrite -eq_rect_eq. -Qed. - -Definition cputchk T (r : loc T) (s : coq_type T) : - cput r s >> cchk r = cput r s. -Proof. -congr mkActo. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -rewrite /coerce. -case: ml_type_eq_dec => HT //. -subst T'. -rewrite -!eq_rect_eq. -do! rewrite bindE /=. -rewrite nth_error_set_nth. -by case: ml_type_eq_dec. -Qed. - -Definition cputput T (r : loc T) (s s' : coq_type T) : - cput r s >> cput r s' = cput r s'. -Proof. -congr mkActo. -apply/boolp.funext => st. -case: r s s' => {}T n s s' /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T'. -rewrite !bindretf /= nth_error_set_nth /coerce. -case: ml_type_eq_dec => // H. -rewrite -eq_rect_eq. -by rewrite set_set_nth eqxx. -Qed. - -Definition cputget T (r : loc T) (s : coq_type T) (A : UU0) (k : coq_type T -> M A) : - cput r s >> (cget r >>= k) = cput r s >> k s. -Proof. -congr mkActo. -apply/boolp.funext => st /=. -case: r s k => {}T n s k /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T'. -rewrite !bindretf /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite nth_error_set_nth. -rewrite /coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -by rewrite -eq_rect_eq !bindretf. -Qed. - -Definition cgetputchk T (r : loc T) : cget r >>= cput r = cchk r. -Proof. -congr mkActo. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite /coerce. -case: (ml_type_eq_dec T1 T) => H //. -subst T1. -rewrite bindE /= bindE /= Hr. -case: (ml_type_eq_dec T T) => H //. -by rewrite -eq_rect_eq nth_error_set_nth_id. -Qed. - -Definition cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : - cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. -Proof. -congr mkActo. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite /coerce. -case: (ml_type_eq_dec T1 T) => H //. -subst T1. -rewrite bindE /= bindE /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite Hr. -case: (ml_type_eq_dec T T) => H //. -by rewrite -eq_rect_eq. -Qed. - -Definition cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) - (k : coq_type T1 -> coq_type T2 -> M A) : - cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = - cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). -Proof. -congr mkActo. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hr1: (nth_error _ _) => [[T1' u]|]; last first. - rewrite bindfailf. - case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. - case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. - rewrite !bindE /=. - by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. -case Hr2: (nth_error _ _) => [[T2' v]|]; last first. - rewrite bindfailf. - case Hc: (coerce T1 u) => [v|]; last by rewrite bindfailf. - rewrite !bindE /=. - by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr2 bindfailf. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec _ _) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE /=. - rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr1 {1}/coerce. - case: (ml_type_eq_dec _ _) => //; by rewrite bindfailf. -subst T1'. -rewrite -eq_rect_eq. -rewrite bindE /=. -rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr2 {1}/coerce. -case: (ml_type_eq_dec T2' T2) => HT2; last by rewrite !bindfailf. -subst T2'. -rewrite -eq_rect_eq. -rewrite !bindE /=. -rewrite !bindE /=. -rewrite /bindS MS_mapE /= fmapE /= bindA /= Hr1 /coerce. -case: (ml_type_eq_dec T1 T1) => // H. -by rewrite -eq_rect_eq. -Qed. - -Definition cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) +Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) (s2 : coq_type T2) (A : UU0) : loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. @@ -643,7 +481,7 @@ move=> H. congr mkActo. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. case Hr1: (nth_error _ _) => [[T1' u]|]; last first. rewrite bindfailf. case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. @@ -702,7 +540,7 @@ case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq set_set_nth (negbTE Hr). Qed. -Definition cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) +Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) (A : UU0) (k : coq_type T2 -> M A) : loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). @@ -712,7 +550,7 @@ congr mkActo. apply/boolp.funext => st /=. rewrite !bindA. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. case Hr1: (nth_error _ _) => [[T1' u]|]; last first. rewrite bindfailf. case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. @@ -743,55 +581,52 @@ case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq bindE. Qed. -Definition crunret (A B : UU0) (m : M A) (s : B) : - crun m -> crun (m >> ret s) = Some s. +Let cputnewC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A + (k : loc T' -> M A) : + cget r >> (cnew s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew s' >>= k). +Proof. +congr mkActo. +apply/boolp.funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. +subst T1. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite (nth_error_rcons_some _ Hr). +rewrite /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /=. +rewrite (nth_error_size_set_nth _ _ Hr). +by rewrite (nth_error_set_nth_rcons _ _ _ Hr). +Qed. + +Let crunret (A B : UU0) (m : M A) (s : B) : + crun m -> crun (m >> Ret s) = Some s. Proof. -case: m => T /= m. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -by case Hm: (m [::]). +by case Hm: (ofActo m [::]). Qed. -Definition crunskip : crun skip = Some tt. +Let crunskip : crun skip = Some tt. Proof. by []. Qed. -Definition crunnew (A : UU0) T (m : M A) (s : coq_type T) : +Let crunnew (A : UU0) T (m : M A) (s : coq_type T) : crun m -> crun (m >> cnew s). Proof. -case: m => U /= m. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -by case Hm: (m [::]). +by case Hm: (ofActo m [::]). Qed. Unset Universe Checking. -Definition isMonadTypedStoreModel : - isMonadTypedStore.axioms_ M isFunctorTS isMonadTS. -Proof. -exists cnew cget cput cchk crun. -- exact cnewget. -- exact cnewgetC. -- exact cnewput. -- exact cnewputC. -- exact cputput. -- exact cputget. -- exact cgetputchk. -- exact cgetget. -- exact cgetC. -- exact cputC. -- exact cputgetC. -- exact cnewchk. -- exact cchknewC. -- exact cchkgetC. -- exact cchknew. -- exact cchkget. -- exact cgetchk. -- exact cchkputC. -- exact cchkput. -- exact cputchk. -- exact cchkC. -- exact cchkdup. -- exact crunret. -- exact crunskip. -- exact crunnew. -Defined. +Canonical Structure isMonadTypedStoreModel := + isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip + cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC + cputgetC cputnewC crunret crunskip crunnew. Set Universe Checking. End ModelTypedStore. From af3ca6271e9ebe2650c5c2635ee015362111f8e6 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 10 Feb 2023 14:46:57 +0900 Subject: [PATCH 22/82] wip : for loop example --- example_typed_store.v | 52 +++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 52 insertions(+) diff --git a/example_typed_store.v b/example_typed_store.v index 52c28312..189871ad 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -90,6 +90,40 @@ by rewrite cnewput IH // (mulnC m.+1) -mulnA. Qed. End factorial. +Section fact_for. +Variable M : typedStoreMonad. +Notation coq_type := (@MLtypes.coq_type M). + +Definition forloop (n_1 n_2 : nat) (b : nat -> M unit) : M unit := + if n_2 < n_1 then Ret tt else + iteri (n_2.+1 - n_1) + (fun i (m : M unit) => m >> b (n_1 + i)) + skip. + +Notation "'do' x <- m ; e" := (m >>= (fun x => e)) + (at level 60, x name, m at level 200, e at level 60). + +Definition fact_for (n : coq_type ml_int) : M (coq_type ml_int) := + do v <- cnew ml_int 1; + do _ <- + (do u <- Ret 1; + do v_1 <- Ret n; + forloop u v_1 + (fun i => + do v_1 <- (do v_1 <- cget v; Ret (v_1 * i)); + cput v v_1)); + cget v. + +Theorem fact_for_ok n : crun (fact_for n) = Some (fact_rec n). +Proof. +rewrite /fact_for /forloop. +under eq_bind do rewrite !bindA !bindretf. +elim: n => [|n IH] /=. + under eq_bind do rewrite bindretf -cgetret. + by rewrite cnewget crunret // crunnew0. +Abort. +End fact_for. + Section fibonacci. Variable M : typedStoreMonad. Notation coq_type := (@MLtypes.coq_type M). @@ -146,3 +180,21 @@ rewrite cnewput. by under eq_bind do rewrite cnewput. Qed. End fibonacci. + +(* +Require Import PrimInt63. +Require Sint63. + +Section fact_for. +Variable M : typedStoreMonad. +Notation coq_type := (@MLtypes.coq_type M). + +Definition nat_of_uint (n : int) : nat := + if Uint63.to_Z n is Zpos pos then Pos.to_nat pos else 0. + +Definition forloop (n_1 n_2 : int) (b : int -> M unit) : M unit := + if Sint63.ltb n_2 n_1 then Ret tt else + iteri (nat_of_uint (sub n_2 n_1)) + (fun i (m : M unit) => m >> b (add n_1 (Uint63.of_Z (Z.of_nat i)))) + (Ret tt). +*) From b686852ee24c1755be465251ca91d0da28c3b51e Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 10 Feb 2023 17:21:30 +0900 Subject: [PATCH 23/82] prove fact_for_ok --- example_typed_store.v | 47 ++++++++++++++++++++++++++++++++++++------- 1 file changed, 40 insertions(+), 7 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 189871ad..ce6b79ed 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -114,14 +114,48 @@ Definition fact_for (n : coq_type ml_int) : M (coq_type ml_int) := cput v v_1)); cget v. +Lemma iteriSr T n (f : nat -> T -> T) x : + iteri n.+1 f x = iteri n (f \o succn) (f 0 x). +Proof. by elim: n x => // n IH x /=; rewrite -IH. Qed. + +Lemma iteri_bind (S : monad) n (f : nat -> S unit) (m1 m2 : S unit) : + iteri n (fun i (m : S unit) => m >> f i) (m1 >> m2) = + m1 >> iteri n (fun i (m : S unit) => m >> f i) m2. +Proof. by elim: n m2 => // n IH m2; rewrite iteriS IH !bindA. Qed. + Theorem fact_for_ok n : crun (fact_for n) = Some (fact_rec n). Proof. rewrite /fact_for /forloop. under eq_bind do rewrite !bindA !bindretf. -elim: n => [|n IH] /=. +transitivity (crun (cnew ml_int (fact_rec n) >> Ret (fact_rec n) : M _)); + last by rewrite crunret // crunnew0. +congr crun. +rewrite subn1 succnK. +pose i := 0. +rewrite -{1}/(fact_rec 0) -{1 3}/i. +set m := n - i. +have {1 2}-> : n = m by rewrite /m subn0. +have -> : n = m + i by rewrite /m subn0 addn0. +clearbody m. +elim: m i {n} => [|m IH] i. under eq_bind do rewrite bindretf -cgetret. - by rewrite cnewget crunret // crunnew0. -Abort. + by rewrite cnewget. +under eq_bind do + rewrite iteriSr bindskipf -[X in iteri _ _ X]bindmskip iteri_bind. +rewrite /=. +case Hm: (m < 1) IH => IH. + case: m Hm {IH} => //= _. + under eq_bind do rewrite !bindA bindretf -cgetret !bindA. + rewrite cnewget. + under eq_bind do rewrite !bindretf. + by rewrite cnewput cnewget addn0 (mulnC i.+1). +under eq_bind do rewrite !bindA. +rewrite cnewget. +under eq_bind do rewrite bindretf. +rewrite cnewput addn0. +under eq_bind do under [fun i : nat => _]boolp.funext do rewrite -addSnnS. +by rewrite mulnC -/(fact_rec i.+1) IH addnS. +Qed. End fact_for. Section fibonacci. @@ -181,20 +215,19 @@ by under eq_bind do rewrite cnewput. Qed. End fibonacci. -(* Require Import PrimInt63. Require Sint63. -Section fact_for. +Section fact_for_int63. Variable M : typedStoreMonad. Notation coq_type := (@MLtypes.coq_type M). Definition nat_of_uint (n : int) : nat := if Uint63.to_Z n is Zpos pos then Pos.to_nat pos else 0. -Definition forloop (n_1 n_2 : int) (b : int -> M unit) : M unit := +Definition forloop63 (n_1 n_2 : int) (b : int -> M unit) : M unit := if Sint63.ltb n_2 n_1 then Ret tt else iteri (nat_of_uint (sub n_2 n_1)) (fun i (m : M unit) => m >> b (add n_1 (Uint63.of_Z (Z.of_nat i)))) (Ret tt). -*) +End fact_for_int63. From 844eff0e8b4857c18557ef1906dab2709c7e146c Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 10 Feb 2023 21:51:22 +0900 Subject: [PATCH 24/82] prove equations for forloop --- example_typed_store.v | 83 ++++++++++++++++++++++++++----------------- 1 file changed, 50 insertions(+), 33 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index ce6b79ed..1ae440ac 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -90,9 +90,8 @@ by rewrite cnewput IH // (mulnC m.+1) -mulnA. Qed. End factorial. -Section fact_for. -Variable M : typedStoreMonad. -Notation coq_type := (@MLtypes.coq_type M). +Section forloop. +Variable M : monad. Definition forloop (n_1 n_2 : nat) (b : nat -> M unit) : M unit := if n_2 < n_1 then Ret tt else @@ -100,6 +99,40 @@ Definition forloop (n_1 n_2 : nat) (b : nat -> M unit) : M unit := (fun i (m : M unit) => m >> b (n_1 + i)) skip. +Lemma iteriSr T n (f : nat -> T -> T) x : + iteri n.+1 f x = iteri n (f \o succn) (f 0 x). +Proof. by elim: n x => // n IH x /=; rewrite -IH. Qed. + +Lemma iteriD T n m (f : nat -> T -> T) x : + iteri (n + m) f x = iteri m (f \o addn n) (iteri n f x). +Proof. by elim: n x f => // n IH x f; rewrite addSn iteriSr IH iteriSr. Qed. + +Lemma iteri_bind n (f : nat -> M unit) (m1 m2 : M unit) : + iteri n (fun i (m : M unit) => m >> f i) (m1 >> m2) = + m1 >> iteri n (fun i (m : M unit) => m >> f i) m2. +Proof. by elim: n m2 => // n IH m2; rewrite iteriS IH !bindA. Qed. + +Lemma forloopS m n (f : nat -> M unit) : + m <= n -> forloop m n f = f m >> forloop m.+1 n f. +Proof. +rewrite /forloop => mn. +rewrite ltnNge mn /= subSS subSn // iteriSr bindskipf. +rewrite -[f _]bindmskip iteri_bind addn0 ltnS -subn_eq0. +case: (n-m) => //= k. +rewrite addSnnS; apply eq_bind => _; congr bind. +apply eq_iteri => i x; by rewrite addSnnS. +Qed. + +Lemma forloop0 m n (f : nat -> M unit) : + m > n -> forloop m n f = skip. +Proof. by rewrite /forloop => ->. Qed. +End forloop. +Arguments forloop {M}. + +Section fact_for. +Variable M : typedStoreMonad. +Notation coq_type := (@MLtypes.coq_type M). + Notation "'do' x <- m ; e" := (m >>= (fun x => e)) (at level 60, x name, m at level 200, e at level 60). @@ -114,47 +147,31 @@ Definition fact_for (n : coq_type ml_int) : M (coq_type ml_int) := cput v v_1)); cget v. -Lemma iteriSr T n (f : nat -> T -> T) x : - iteri n.+1 f x = iteri n (f \o succn) (f 0 x). -Proof. by elim: n x => // n IH x /=; rewrite -IH. Qed. - -Lemma iteri_bind (S : monad) n (f : nat -> S unit) (m1 m2 : S unit) : - iteri n (fun i (m : S unit) => m >> f i) (m1 >> m2) = - m1 >> iteri n (fun i (m : S unit) => m >> f i) m2. -Proof. by elim: n m2 => // n IH m2; rewrite iteriS IH !bindA. Qed. - Theorem fact_for_ok n : crun (fact_for n) = Some (fact_rec n). Proof. -rewrite /fact_for /forloop. +rewrite /fact_for. under eq_bind do rewrite !bindA !bindretf. transitivity (crun (cnew ml_int (fact_rec n) >> Ret (fact_rec n) : M _)); last by rewrite crunret // crunnew0. congr crun. -rewrite subn1 succnK. -pose i := 0. -rewrite -{1}/(fact_rec 0) -{1 3}/i. -set m := n - i. -have {1 2}-> : n = m by rewrite /m subn0. -have -> : n = m + i by rewrite /m subn0 addn0. +rewrite -{1}/(fact_rec 0). +pose m := n. +have -> : 0 = n - m by rewrite subnn. +have : m <= n by []. clearbody m. -elim: m i {n} => [|m IH] i. - under eq_bind do rewrite bindretf -cgetret. +elim: m => [|m IH] mn. + rewrite subn0. + under eq_bind do rewrite forloop0 ?leqnn // bindretf -cgetret. by rewrite cnewget. -under eq_bind do - rewrite iteriSr bindskipf -[X in iteri _ _ X]bindmskip iteri_bind. -rewrite /=. -case Hm: (m < 1) IH => IH. - case: m Hm {IH} => //= _. - under eq_bind do rewrite !bindA bindretf -cgetret !bindA. - rewrite cnewget. - under eq_bind do rewrite !bindretf. - by rewrite cnewput cnewget addn0 (mulnC i.+1). +rewrite subnSK //. +under eq_bind do (rewrite forloopS; last by apply leq_subr). under eq_bind do rewrite !bindA. rewrite cnewget. under eq_bind do rewrite bindretf. -rewrite cnewput addn0. -under eq_bind do under [fun i : nat => _]boolp.funext do rewrite -addSnnS. -by rewrite mulnC -/(fact_rec i.+1) IH addnS. +rewrite cnewput. +rewrite subnS (_ : fact_rec _ * _ = fact_rec (n - m)). + by rewrite IH // ltnW. +by rewrite mulnC -(@prednK (n-m)) // lt0n subn_eq0 -ltnNge. Qed. End fact_for. From fb5df79be4c1af88d959f4551d53ce93104ac877 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 10 Feb 2023 23:07:47 +0900 Subject: [PATCH 25/82] simplify and update impredicative_set --- example_typed_store.v | 7 +- impredicative_set/ihierarchy.v | 151 ++++++-- impredicative_set/imonad_model.v | 587 ++++++++++++++++++++++++++----- 3 files changed, 621 insertions(+), 124 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 1ae440ac..61a01edd 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -158,7 +158,6 @@ rewrite -{1}/(fact_rec 0). pose m := n. have -> : 0 = n - m by rewrite subnn. have : m <= n by []. -clearbody m. elim: m => [|m IH] mn. rewrite subn0. under eq_bind do rewrite forloop0 ?leqnn // bindretf -cgetret. @@ -168,10 +167,8 @@ under eq_bind do (rewrite forloopS; last by apply leq_subr). under eq_bind do rewrite !bindA. rewrite cnewget. under eq_bind do rewrite bindretf. -rewrite cnewput. -rewrite subnS (_ : fact_rec _ * _ = fact_rec (n - m)). - by rewrite IH // ltnW. -by rewrite mulnC -(@prednK (n-m)) // lt0n subn_eq0 -ltnNge. +rewrite cnewput -IH; last by apply ltnW. +by rewrite subnS mulnC -(@prednK (n-m)) // lt0n subn_eq0 -ltnNge. Qed. End fact_for. diff --git a/impredicative_set/ihierarchy.v b/impredicative_set/ihierarchy.v index 8c3b99b4..a8722b2c 100644 --- a/impredicative_set/ihierarchy.v +++ b/impredicative_set/ihierarchy.v @@ -2,7 +2,7 @@ (* Copyright (C) 2020 monae authors, license: LGPL-2.1-or-later *) Ltac typeof X := type of X. -Require Import ssrmatching. +Require Import ssrmatching JMeq. From mathcomp Require Import all_ssreflect. Require Import imonae_lib. From HB Require Import structures. @@ -1069,43 +1069,56 @@ HB.structure Definition MonadPlusArray (S : UU0) (I : eqType) := Module Type MLTY. Parameter ml_type : UU0. Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Variant loc : ml_type -> UU0 := mkloc T : nat -> loc T. +Parameter loc : ml_type -> UU0. +Parameter locT : eqType. +Parameter loc_id : forall T, loc T -> locT. Parameter coq_type : forall M : UU0 -> UU0, ml_type -> UU0. End MLTY. Module MonadTypedStore (MLtypes : MLTY). Import MLtypes. -Record binding (M : UU0 -> UU0) := - mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. -Arguments mkbind {M}. - -Definition loc_id {T} (r : loc T) : nat := let: mkloc _ n := r in n. - HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) of Monad M := { cnew : forall {T}, coq_type M T -> M (loc T) ; cget : forall {T}, loc T -> M (coq_type M T) ; cput : forall {T}, loc T -> coq_type M T -> M unit ; - cchk : forall {T}, loc T -> M unit ; - cputput : forall T (r : loc T) (s s' : coq_type M T), - cput r s >> cput r s' = cput r s' ; - cputget : - forall T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A), - cput r s >> cget r >>= k = cput r s >> k s ; - cgetputchk : forall T (r : loc T), cget r >>= cput r = cchk r ; + crun : forall {A : UU0}, M A -> option A ; (* execute in empty store *) + cnewget : forall T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A), + cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s) ; + cnewput : forall T (s t : coq_type M T) A (k : loc T -> M A), + cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k ; + cgetput : forall T (r : loc T) (s : coq_type M T), + cget r >> cput r s = cput r s ; + cgetputskip : forall T (r : loc T), cget r >>= cput r = cget r >> skip ; cgetget : forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s ; + cputget : + forall T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A), + cput r s >> (cget r >>= k) = cput r s >> k s ; + cputput : forall T (r : loc T) (s s' : coq_type M T), + cput r s >> cput r s' = cput r s' ; cgetC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) (k : coq_type M T1 -> coq_type M T2 -> M A), cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)) ; + cgetnewD : + forall T T' (r : loc T) (s : coq_type M T') A + (k : loc T' -> coq_type M T -> coq_type M T -> M A), + cget r >>= (fun u => cnew s >>= (fun r' => cget r >>= k r' u)) = + cget r >>= (fun u => cnew s >>= (fun r' => k r' u u)) ; + cgetnewE : forall T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) + (k1 k2 : loc T2 -> M A), + (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2) ; + cgetputC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2), + cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip ; cputC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) (s2 : coq_type M T2) (A : UU0), - loc_id r1 != loc_id r2 \/ mkbind T1 s1 = mkbind T2 s2 -> + loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1 ; cputgetC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) @@ -1113,18 +1126,110 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v) ; - cnewget : forall T (s : coq_type M T) A (k : coq_type M T -> M A), - cnew s >>= (fun r => cget r >>= k) = cnew s >> k s ; - cnewgetC : - forall T T' (r : loc T) (s : coq_type M T') A - (k : loc T' -> coq_type M T -> M A), - cnew s >>= (fun r' => cget r >>= k r') = - cget r >>= (fun u => cnew s >>= (fun r' => k r' u)) ; + cputnewC : + forall T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A + (k : loc T' -> M A), + cget r >> (cnew s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew s' >>= k) ; + crunret : forall (A B : UU0) (m : M A) (s : B), + crun m -> crun (m >> Ret s) = Some s ; + crunskip : + crun skip = Some tt ; + crunnew : forall (A : UU0) T (m : M A) (s : coq_type M T), + crun m -> crun (m >> cnew s) ; }. #[short(type=typedStoreMonad)] HB.structure Definition MonadTypedStore := { M of isMonadTypedStore M & isMonad M & isFunctor M }. + +Arguments cnew {s}. +Arguments cget {s} [T]. +Arguments cput {s} [T]. +Arguments crun {s} [A]. + +Section cchk. +Variable M : typedStoreMonad. +Definition cchk T (r : loc T) : M unit := cget r >> skip. + +Lemma cnewchk T s (A : UU0) (k : loc T -> M A) : + cnew T s >>= (fun r => cchk r >> k r) = cnew T s >>= k. +Proof. +under eq_bind do rewrite bindA. +rewrite cnewget. +by under eq_bind do rewrite bindskipf. +Qed. + +Lemma cchknewC T1 T2 (r : loc T1) s (A : UU0) (k : loc T2 -> M A) : + cchk r >> (cnew T2 s >>= fun r' => cchk r >> k r') = + cchk r >> (cnew T2 s >>= k). +Proof. +rewrite !(bindA,bindskipf). +under eq_bind do under eq_bind do rewrite bindA. +rewrite cgetnewD. +by under eq_bind do under eq_bind do rewrite bindskipf. +Qed. + +Lemma cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) + (k: coq_type M T2 -> M A) : + cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). +Proof. +under [in RHS]eq_bind do rewrite bindA bindskipf. +by rewrite !(bindA,bindskipf) cgetC. +Qed. + +Lemma cchknewE T1 T2 (r1 : loc T1) s (A : UU0) (k1 k2 : loc T2 -> M A) : + (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cchk r1 >> (cnew T2 s >>= k1) = cchk r1 >> (cnew T2 s >>= k2). +Proof. move=> Hk; rewrite !(bindA,bindskipf); exact: cgetnewE. Qed. + +Lemma cchknewget T T' (r : loc T) s (A : UU0) k : + cchk r >> (cnew T' s >>= fun r' => cget r >>= k r') = + cget r >>= (fun u => cnew T' s >>= k ^~ u) :> M A. +Proof. by rewrite bindA bindskipf cgetnewD. Qed. + +Lemma cchknewput T T' (r : loc T) s s' (A : UU0) k : + cchk r >> (cnew T' s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew T' s' >>= k) :> M A. +Proof. by rewrite bindA bindskipf cputnewC. Qed. + +Lemma cchkget T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : + cchk r >> (cget r >>= k) = cget r >>= k. +Proof. by rewrite bindA bindskipf cgetget. Qed. + +Lemma cgetchk T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : + cget r >>= (fun s => cchk r >> k s) = cget r >>= k. +Proof. under eq_bind do rewrite bindA bindskipf; by rewrite cgetget. Qed. + +Lemma cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : + cchk r1 >> cput r2 s = cput r2 s >> cchk r1. +Proof. by rewrite bindA bindskipf cgetputC bindA. Qed. + +Lemma cchkput T (r : loc T) (s : coq_type M T) : + cchk r >> cput r s = cput r s. +Proof. by rewrite bindA bindskipf cgetput. Qed. + +Lemma cputchk T (r : loc T) (s : coq_type M T) : + cput r s >> cchk r = cput r s. +Proof. by rewrite cputget bindmskip. Qed. + +Lemma cchkC T1 T2 (r1: loc T1) (r2: loc T2) : + cchk r1 >> cchk r2 = cchk r2 >> cchk r1. +Proof. by rewrite !(bindA,bindskipf) cgetC. Qed. + +Lemma cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. +Proof. by rewrite bindA bindskipf cgetget. Qed. + +Lemma cgetret T (r : loc T) : cget r >>= Ret = cget r :> M _. +Proof. by rewrite bindmret. Qed. + +Lemma crunnew0 T s : crun (cnew T s : M _). +Proof. by rewrite -(bindskipf (cnew T s)) crunnew // crunskip. Qed. + +Lemma cnewgetret T s : cnew T s >>= @cget _ _ = cnew T s >> Ret s :> M _. +Proof. under eq_bind do rewrite -cgetret; by rewrite cnewget. Qed. +End cchk. +Arguments cchk {M} [T]. End MonadTypedStore. HB.mixin Record isMonadTrace (T : UU0) (M : UU0 -> UU0) of Monad M := diff --git a/impredicative_set/imonad_model.v b/impredicative_set/imonad_model.v index 1d8198f2..ddbfba62 100644 --- a/impredicative_set/imonad_model.v +++ b/impredicative_set/imonad_model.v @@ -1118,75 +1118,135 @@ HB.instance Definition _ := isMonadArray.Build End modelarray. End ModelArray. -Module ModelTypedStore (MLtypes : MLTY). -Module MTypedStore := MonadTypedStore (MLtypes). +Section nth_error. +Context [T : UU0] (def : T) (st : seq T). + +(* Basic lemmas for standard library's nth_error *) +Local Notation nth_error := List.nth_error. + +Lemma nth_error_set_nth n x : + nth_error (set_nth def st n x) n = Some x. +Proof. +elim: n st => [|z IH] [] // {IH}. +elim: z.+1 => [|n <-] //=. +by rewrite set_nth_nil. +Qed. + +Lemma nth_error_rcons_size b : + nth_error (rcons st b) (size st) = Some b. +Proof. by elim: st. Qed. + +Lemma nth_error_rcons_some n a b : + nth_error st n = Some a -> nth_error (rcons st b) n = Some a. +Proof. by elim: n st => [|n IH] []. Qed. + +Lemma nth_error_set_nth_id n a : + nth_error st n = Some a -> set_nth def st n a = st. +Proof. by elim: n st => [|n IH] [] //= b st'; [case=> -> | move/IH ->]. Qed. + +Lemma nth_error_set_nth_other m n a b : + m != n -> + nth_error st m = Some a -> + nth_error (set_nth def st n b) m = Some a. +Proof. elim: m st n => [|m IH] [|c st'] [|n] //=; rewrite eqSS => *; exact: IH. +Qed. + +Lemma nth_error_set_nth_none m n a b : + nth_error st m = None -> + nth_error st n = Some a -> + nth_error (set_nth def st n b) m = None. +Proof. by elim: m st n => [|m IH] [|c st'] [|n] //=; apply IH. Qed. + +Lemma nth_error_size n a : + nth_error st n = Some a -> n < size st. +Proof. by elim: n st => [|n IH] [|c st'] //= /IH. Qed. + +Lemma nth_error_size_set_nth n a b : + nth_error st n = Some a -> size (set_nth def st n b) = size st. +Proof. by rewrite size_set_nth => /nth_error_size /maxn_idPr. Qed. + +Lemma set_nth_rcons a b : + set_nth def (rcons st a) (size st) b = rcons st b. +Proof. by elim: st => //= c st' ->. Qed. + +Lemma nth_error_set_nth_rcons n a b c : + nth_error st n = Some a -> + set_nth def (rcons st c) n b = rcons (set_nth def st n b) c. +Proof. by elim: n st => [|n IH] [|d st'] //= /IH ->. Qed. +End nth_error. + +Require Import JMeq. + +Module Type MLTYweak. +Parameter ml_type : Set. +Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. +Variant loc : ml_type -> Set := mkloc T : nat -> loc T. +Parameter coq_type0 : ml_type -> UU0. +Parameter ml_undef : ml_type. +Parameter undef : coq_type0 ml_undef. +End MLTYweak. + +Module ModelTypedStore (MLtypes : MLTYweak). +Import MLtypes. +Module MLtypes'. +Definition ml_type := ml_type. +Definition ml_type_eq_dec := ml_type_eq_dec. +Definition coq_type (M : UU0 -> UU0) := coq_type0. +Definition loc := loc. +Definition locT := [eqType of nat]. +Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. +End MLtypes'. +Module MTypedStore := MonadTypedStore (MLtypes'). +Import MLtypes'. Import MLtypes. Import MTypedStore. -#[bypass_check(positivity)] -Inductive acto : UU0 -> UU0 := - mkActo : forall T : UU0, MS (seq (binding acto)) [the monad of option_monad] T -> acto T. +Record binding (M : UU0 -> UU0) := + mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. +Arguments mkbind {M}. + +Section typed_store. +Definition acto (T : UU0) : UU0 := + MS (seq (binding idfun)) [the monad of option_monad] T. Local Notation M := acto. -Local Notation coq_type := (coq_type M). +(*Local Notation coq_type := (coq_type M).*) +Local Notation coq_type' := (coq_type idfun). -Definition ofActo (T:UU0) (m : M T) : MS (seq (binding M)) [the monad of option_monad] T := - let: mkActo _ m' := m in m'. +Definition def : binding idfun := mkbind ml_undef undef. +Local Notation nth_error := List.nth_error. -Definition cnew T (v : coq_type T) : M (loc T) := - mkActo (fun st => let n := size st in Ret (mkloc T n, rcons st (mkbind T v))). +Definition cnew T (v : coq_type M T) : M (loc T) := + fun st => let n := size st in + Ret (mkloc T n, rcons st (mkbind T (v : coq_type' T))). -Definition coerce (T1 T2 : ml_type) (v : coq_type T1) : option (coq_type T2) := +Definition coerce T1 T2 (v : coq_type M T1) : option (coq_type M T2) := match ml_type_eq_dec T1 T2 with | left H => Some (eq_rect _ _ v _ H) | right _ => None end. -Definition cget T (r : loc T) : M (coq_type T) := - mkActo (fun st => - if List.nth_error st (loc_id r) is Some (mkbind T' v) then - if coerce T v is Some u then Ret (u, st) else fail - else fail). - -Definition cput T (r : loc T) (v : coq_type T) : M unit := - mkActo (fun st => - let n := loc_id r in - if List.nth_error st n is Some (mkbind T' _) then - if coerce T' v is Some u then - let b := mkbind T' u in - Ret (tt, set_nth b st n b) - else fail - else fail). - -Definition cchk T (r : loc T) : M unit := - mkActo (fun st => - if List.nth_error st (loc_id r) is Some (mkbind T' u) then - if coerce T u is Some _ then Ret (tt, st) else fail - else fail). - -Let ret : forall A : UU0, idfun A -> M A := fun A a => mkActo (Ret a). -Let bind (A B : UU0) (m : M A) (f : A -> M B) : M B := - mkActo (ofActo m >>= (fun a => ofActo (f a))). -Let left_neutral : BindLaws.left_neutral bind ret. -Proof. by move=> A B a f; rewrite /bind /ret bindretf; case: (f a). Qed. -Let right_neutral : BindLaws.right_neutral bind ret. -Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. -Let associative : BindLaws.associative bind. -Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. - -(* -HB.instance Definition xyz := - isMonad_ret_bind.Build M left_neutral right_neutral associative. -*) - -Lemma nth_error_set_nth T (x : T) st n : - List.nth_error (set_nth x st n x) n = Some x. -Proof. -elim: n st => [|z IH] [] //. -clear IH. -elim: z.+1 => [|n <-] //=. -by rewrite set_nth_nil. -Qed. +Definition cget T (r : loc T) : M (coq_type M T) := + fun st => + if nth_error st (loc_id r) is Some (mkbind T' v) then + if coerce T v is Some u then Ret (u, st) else fail + else fail. + +Definition cput T (r : loc T) (v : coq_type M T) : M unit := + fun st => + let n := loc_id r in + if nth_error st n is Some (mkbind T' _) then + if coerce T' v is Some u then + Ret (tt, set_nth def st n (mkbind T' (u : coq_type' _))) + else fail + else fail. + +Definition crun (A : UU0) (m : M A) : option A := + match m nil with + | inl _ => None + | inr (a, _) => Some a + end. +(* Make ml_type an eqType *) Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. Lemma ml_type_eqP : Equality.axiom ml_type_eqb. Proof. @@ -1196,34 +1256,90 @@ Qed. Definition ml_type_eq_mixin := EqMixin ml_type_eqP. Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. -Let cputput T (r : loc T) (s s' : coq_type T) : - bind (cput r s) (fun _ => cput r s') = cput r s'. +(* Prove the laws *) +Let cnewget T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A) : + cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. -congr mkActo. -apply/funext => st. -case: r s s' => {}T n s s' /=. +apply/funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hst : (List.nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T'. -rewrite !bindretf /= nth_error_set_nth /coerce. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite /cget nth_error_rcons_size /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cnewput T (s t : coq_type M T) A (k : loc T -> M A) : + cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. +Proof. +apply/funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE bindA. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= [in RHS]bindE /= /cput. +rewrite nth_error_rcons_size. +rewrite /coerce. case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. +Qed. + +Let cgetput T (r : loc T) (s : coq_type M T) : + cget r >> cput r s = cput r s. +Proof. +apply/funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. +case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. +rewrite /coerce. +case: ml_type_eq_dec => HT; last by case: ml_type_eq_dec => // /esym //. +subst T'. rewrite -eq_rect_eq. -set b := mkbind _ _. -set b' := mkbind _ _. -admit. -Admitted. +case: ml_type_eq_dec => // HT. +rewrite -eq_rect_eq {HT}. +do! rewrite bindE /=. +rewrite Hr. +case: ml_type_eq_dec => // HT. +by rewrite -eq_rect_eq. +Qed. -Let cputget T (r : loc T) (s : coq_type T) (A : UU0) (k : coq_type T -> M A) : - bind (cput r s) (fun=> bind (cget r) k) = bind (cput r s) (fun=> k s). +Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. +Proof. +apply/funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE bindA /=. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite /coerce. +case: (ml_type_eq_dec T1 T) => H //. +subst T1. +rewrite bindE /= bindE /= Hr. +case: (ml_type_eq_dec T T) => H //. +by rewrite -eq_rect_eq nth_error_set_nth_id. +Qed. + +Let cgetget T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A) : + cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. +Proof. +apply/funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite /coerce. +case: (ml_type_eq_dec T1 T) => H //. +subst T1. +rewrite bindE /= bindE /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite Hr. +case: (ml_type_eq_dec T T) => H //. +by rewrite -eq_rect_eq. +Qed. + +Let cputget T (r: loc T) (s: coq_type M T) (A: UU0) (k: coq_type M T -> M A) : + cput r s >> (cget r >>= k) = cput r s >> k s. Proof. -congr mkActo. apply/funext => st /=. case: r s k => {}T n s k /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -case Hst : (List.nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. +case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. rewrite {1 3}/coerce. case: ml_type_eq_dec => H /=; last by rewrite bindfailf. subst T'. @@ -1235,28 +1351,307 @@ case: ml_type_eq_dec => H /=; last by rewrite bindfailf. by rewrite -eq_rect_eq !bindretf. Qed. -(* - cgetputchk : forall T (r : loc T), cget r >>= cput r = cchk r ; - cgetget : - forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), - cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s ; - cgetC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) - (k : coq_type M T1 -> coq_type M T2 -> M A), - cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = - cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)) ; - cputC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) - (s2 : coq_type M T2) (A : UU0), - loc_id r1 != loc_id r2 \/ mkbind T1 s1 = mkbind T2 s2 -> - cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1 ; - cputgetC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) - (A : UU0) (k : coq_type M T2 -> M A), - loc_id r1 != loc_id r2 -> - cput r1 s1 >> cget r2 >>= k = - cget r2 >>= (fun v => cput r1 s1 >> k v) }. +Let cputput T (r : loc T) (s s' : coq_type M T) : + cput r s >> cput r s' = cput r s'. +Proof. +apply/funext => st. +case: r s s' => {}T n s s' /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. +case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T'. +rewrite !bindretf /= nth_error_set_nth /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +by rewrite set_set_nth eqxx. +Qed. + +Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) + (k : coq_type M T1 -> coq_type M T2 -> M A) : + cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = + cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). +Proof. +apply/funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. +case Hr1: (nth_error _ _) => [[T1' u]|]; last first. + rewrite bindfailf. + case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. + case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. + rewrite !bindE /=. + by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. +case Hr2: (nth_error _ _) => [[T2' v]|]; last first. + rewrite bindfailf. + case Hc: (coerce T1 u) => [v|]; last by rewrite bindfailf. + rewrite !bindE /=. + by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr2 bindfailf. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec _ _) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE /=. + rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr1 {1}/coerce. + case: (ml_type_eq_dec _ _) => //; by rewrite bindfailf. +subst T1'. +rewrite -eq_rect_eq. +rewrite bindE /=. +rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr2 {1}/coerce. +case: (ml_type_eq_dec T2' T2) => HT2; last by rewrite !bindfailf. +subst T2'. +rewrite -eq_rect_eq. +rewrite !bindE /=. +rewrite !bindE /=. +rewrite /bindS MS_mapE /= fmapE /= bindA /= Hr1 /coerce. +case: (ml_type_eq_dec T1 T1) => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cgetnewD T T' (r : loc T) (s : coq_type M T') A + (k : loc T' -> coq_type M T -> coq_type M T -> M A) : + cget r >>= (fun u => cnew s >>= fun r' => cget r >>= k r' u) = + cget r >>= (fun u => cnew s >>= fun r' => k r' u u). +Proof. +apply/funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by rewrite bindfailf. +subst T1. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite (nth_error_rcons_some _ Hr). +rewrite /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq. +Qed. + +Let cgetnewE T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) + (k1 k2 : loc T2 -> M A) : + (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2). +Proof. +move=> Hk. +apply/funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] //=. +case Hc: (coerce T1 v1) => [u1|] //. +rewrite bindE /= bindE /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= [in RHS]bindE /= [in RHS]bindE /= /bindS MS_mapE /=. +rewrite fmapE /= bindA /= 2!bindE /=. +by rewrite Hk // neq_ltn /= (nth_error_size Hr1). +Qed. + +Let cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : + cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip. +Proof. +apply/funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= 2!bindA /= /cget /cput. +case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. + rewrite bindfailf. + case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; + last by rewrite bindfailf. + case Hc: (coerce T2' s) => [u|]; last by rewrite bindfailf. + rewrite bindE /= /bindS /= bindE /=. + by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). +case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. + rewrite bindfailf. + case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. + by rewrite bindE/= bindE/= Hr2. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec T1' T1) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec T2 T2') => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE/= bindE/=. + case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. + rewrite -Hr nth_error_set_nth /coerce. + move: Hr2; rewrite -Hr Hr1 => -[] HT2. + subst T1'. + by case: ml_type_eq_dec. + rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. + by case: ml_type_eq_dec. +rewrite bindE /= bindE /= Hr2 {1}/coerce. +case: ml_type_eq_dec => HT2 //. +subst T1' T2'. +rewrite -!eq_rect_eq. +rewrite bindE /= bindE /=. +case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. + rewrite -Hr nth_error_set_nth /coerce. + move: Hr2; rewrite -Hr Hr1 => -[] HT2. + subst T2. + by case: ml_type_eq_dec. +rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. +by case: ml_type_eq_dec. +Qed. + +Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (s2 : coq_type M T2) (A : UU0) : + loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> + cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. +Proof. +move=> H. +apply/funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. +case Hr1: (nth_error _ _) => [[T1' u]|]; last first. + rewrite bindfailf. + case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. + case Hc: (coerce T2' s2) => [u|]; last by rewrite bindfailf. + rewrite !bindE /=. + rewrite !bindE /=. + case/boolP: (loc_id r1 == loc_id r2) => Hr. + by rewrite -(eqP Hr) Hr1 in Hr2. + by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). +case/boolP: (loc_id r1 == loc_id r2) H => [/eqP|] Hr /= H. + rewrite -Hr Hr1. + case/boolP: (T1 == T2) => /eqP HT; last first. + rewrite /coerce. + case: (ml_type_eq_dec T1 T1') => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec T2 T1') => HT2; last by rewrite bindfailf. + subst T1'. + rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. + by case: ml_type_eq_dec. + subst T1'. + rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. + by case: ml_type_eq_dec => [/esym|] //. + subst T2. + rewrite /coerce. + case: ml_type_eq_dec => HT1 //. + subst T1'. + rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /= !nth_error_set_nth. + case: ml_type_eq_dec => HT1 //. + rewrite -!eq_rect_eq. + case: H => // H. + by move: (JMeq_eq H) => <-. +case Hr2: (nth_error _ _) => [[T2' v]|]; last first. + rewrite bindfailf. + case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. + rewrite !bindE /=. + rewrite !bindE /=. + by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1). +rewrite {1 3}/coerce. +case: (ml_type_eq_dec _ _) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE /=. + rewrite !bindE /=. + rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. + by case: ml_type_eq_dec. +subst T1'. +rewrite -eq_rect_eq /coerce. +rewrite bindE /= bindE /= (nth_error_set_nth_other _ _ _ Hr2); + last by rewrite eq_sym. +case: ml_type_eq_dec => HT2' //. +subst T2'. +rewrite -eq_rect_eq bindE /= bindE /= (nth_error_set_nth_other _ _ Hr Hr1). +case: ml_type_eq_dec => // HT1. +by rewrite -eq_rect_eq set_set_nth (negbTE Hr). +Qed. + +Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + (A : UU0) (k : coq_type M T2 -> M A) : + loc_id r1 != loc_id r2 -> + cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). +Proof. +move=> Hr. +apply/funext => st /=. +rewrite !bindA. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. +case Hr1: (nth_error _ _) => [[T1' u]|]; last first. + rewrite bindfailf. + case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. + case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. + by rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. +case Hr2: (nth_error _ _) => [[T2' v]|]; last first. + rewrite bindfailf. + case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. + rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE. + by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1) bindfailf. +rewrite {1 3}/coerce. +case: (ml_type_eq_dec _ _) => HT1; last first. + rewrite bindfailf. + case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + subst T2'. + rewrite -eq_rect_eq. + rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE Hr1 /coerce. + by case: ml_type_eq_dec. +subst T1'. +rewrite -eq_rect_eq /coerce. +rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite (nth_error_set_nth_other _ _ _ Hr2); last by rewrite eq_sym. +case: ml_type_eq_dec => HT2' //. +subst T2'. +rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /=. +rewrite /bindS /= MS_mapE /= fmapE Hr1. +case: ml_type_eq_dec => // HT1. +by rewrite -eq_rect_eq bindE. +Qed. + +Let cputnewC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A + (k : loc T' -> M A) : + cget r >> (cnew s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew s' >>= k). +Proof. +apply/funext => st /=. +rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. +case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +rewrite {1 3}/coerce. +case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. +subst T1. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite (nth_error_rcons_some _ Hr). +rewrite /coerce. +case: ml_type_eq_dec => // H. +rewrite -eq_rect_eq. +rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. +rewrite bindE /= /bindS /=. +rewrite (nth_error_size_set_nth _ _ Hr). +by rewrite (nth_error_set_nth_rcons _ _ _ Hr). +Qed. + +Let crunret (A B : UU0) (m : M A) (s : B) : + crun m -> crun (m >> Ret s) = Some s. +Proof. +rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +by case Hm: (m [::]). +Qed. + +Let crunskip : crun skip = Some tt. +Proof. by []. Qed. + +Let crunnew (A : UU0) T (m : M A) (s : coq_type M T) : + crun m -> crun (m >> cnew s). +Proof. +rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +by case Hm: (m [::]). +Qed. + +Canonical Structure isMonadTypedStoreModel := + isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip + cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC + cputgetC cputnewC crunret crunskip crunnew. + +(* The elpi tactic/command HB.instance failed without giving a specific + error message. +HB.instance Definition _ := + isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip + cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC + cputgetC cputnewC crunret crunskip crunnew. *) +End typed_store. End ModelTypedStore. (* TODO? From 8d91c9b87c396809e69e5ab2bf69912900f65075 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Wed, 22 Feb 2023 19:23:22 +0900 Subject: [PATCH 26/82] change definition forloop63 --- example_typed_store.v | 27 ++++++++++++++++++++++++--- 1 file changed, 24 insertions(+), 3 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 61a01edd..c29c64a8 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -239,9 +239,30 @@ Notation coq_type := (@MLtypes.coq_type M). Definition nat_of_uint (n : int) : nat := if Uint63.to_Z n is Zpos pos then Pos.to_nat pos else 0. +Notation "'do' x <- m ; e" := (m >>= (fun x => e)) + (at level 60, x name, m at level 200, e at level 60). + Definition forloop63 (n_1 n_2 : int) (b : int -> M unit) : M unit := if Sint63.ltb n_2 n_1 then Ret tt else - iteri (nat_of_uint (sub n_2 n_1)) - (fun i (m : M unit) => m >> b (add n_1 (Uint63.of_Z (Z.of_nat i)))) - (Ret tt). + iter (nat_of_uint (sub n_2 n_1)) + (fun (m : M int) => do i <- m; do _ <- b i; Ret (Uint63.succ i)) + (Ret n_1) >> Ret tt. + +Lemma ltsbNlesb m n : ltsb m n = ~~ lesb n m. +Proof. +case/boolP: (lesb n m) => /Sint63.lebP nm; apply/Sint63.ltbP => /=; + by [apply Z.le_ngt | apply Z.nle_gt]. +Qed. + +Lemma forloop63S m n (f : int -> M unit) : + lesb m n -> forloop63 m n f = f m >> forloop63 (add m 1) n f. +Proof. +rewrite /forloop63 => mn. +rewrite ltsbNlesb mn /=. +Abort. + +Lemma forloop630 m n (f : nat -> M unit) : + m > n -> forloop m n f = skip. +Proof. by rewrite /forloop => ->. Qed. + End fact_for_int63. From 4b2a9289204998d5621d2992b830adb6553321da Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 24 Feb 2023 19:02:06 +0900 Subject: [PATCH 27/82] prove 63-bit version (with many admits) --- example_typed_store.v | 117 +++++++++++++++++++++++++++++++++++++++--- 1 file changed, 110 insertions(+), 7 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index c29c64a8..8e8ceca8 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -232,9 +232,38 @@ End fibonacci. Require Import PrimInt63. Require Sint63. +Module MLtypes63. +Local Definition ml_type_eq_dec := ml_type_eq_dec. +Local Definition ml_type := ml_type. +Local Definition undef := Undef. +Local Definition loc := loc. +Local Definition locT := [eqType of nat]. +Local Definition loc_id := @loc_id. + +Section with_monad. +Context [M : Type -> Type]. + +(* Generated type translation function *) +Fixpoint coq_type (T : ml_type) : Type := + match T with + | ml_int => int + | ml_bool => bool + | ml_unit => unit + | ml_arrow T1 T2 => coq_type T1 -> M (coq_type T2) + | ml_ref T1 => loc T1 + | ml_undef => undef_t + end. +End with_monad. +Local Definition ml_undef := ml_undef. +End MLtypes63. + +Module IMonadTS63 := hierarchy.MonadTypedStore (MLtypes63). +Import MLtypes63. +Import IMonadTS63. + Section fact_for_int63. Variable M : typedStoreMonad. -Notation coq_type := (@MLtypes.coq_type M). +Notation coq_type := (@MLtypes63.coq_type M). Definition nat_of_uint (n : int) : nat := if Uint63.to_Z n is Zpos pos then Pos.to_nat pos else 0. @@ -244,7 +273,7 @@ Notation "'do' x <- m ; e" := (m >>= (fun x => e)) Definition forloop63 (n_1 n_2 : int) (b : int -> M unit) : M unit := if Sint63.ltb n_2 n_1 then Ret tt else - iter (nat_of_uint (sub n_2 n_1)) + iter (nat_of_uint (sub n_2 n_1)).+1 (fun (m : M int) => do i <- m; do _ <- b i; Ret (Uint63.succ i)) (Ret n_1) >> Ret tt. @@ -254,15 +283,89 @@ case/boolP: (lesb n m) => /Sint63.lebP nm; apply/Sint63.ltbP => /=; by [apply Z.le_ngt | apply Z.nle_gt]. Qed. +Lemma iter_bind T n (f : T -> M T) (m1 : M unit) m2 : + iter n (fun (m : M T) => m >>= f) (m1 >> m2) = + m1 >> iter n (fun (m : M T) => m >>= f) m2. +Proof. by elim: n m2 => // n IH m2; rewrite iterS IH !bindA. Qed. + +Lemma lesb_ltsb_eq m n : lesb m n -> ltsb n (Uint63.succ m) -> m = n. +Admitted. + +Lemma nat_of_uint_sub_succ m n : lesb m n -> lesb (Uint63.succ m) n -> + nat_of_uint (sub n m) = (nat_of_uint (sub n (Uint63.succ m))).+1. +Admitted. + Lemma forloop63S m n (f : int -> M unit) : - lesb m n -> forloop63 m n f = f m >> forloop63 (add m 1) n f. + lesb m n -> forloop63 m n f = f m >> forloop63 (Uint63.succ m) n f. Proof. rewrite /forloop63 => mn. rewrite ltsbNlesb mn /=. -Abort. +case: ifPn => m1n. + rewrite (lesb_ltsb_eq _ _ mn m1n). + by rewrite Sint63.sub_of_Z Z.sub_diag /= !(bindA,bindretf). +rewrite ltsbNlesb negbK in m1n. +rewrite nat_of_uint_sub_succ //. +by rewrite iterSr bindretf !bindA iter_bind !bindA. +Qed. -Lemma forloop630 m n (f : nat -> M unit) : - m > n -> forloop m n f = skip. -Proof. by rewrite /forloop => ->. Qed. +Lemma forloop630 m n (f : int -> M unit) : + ltsb n m -> forloop63 m n f = skip. +Proof. by rewrite /forloop63 => ->. Qed. + +Definition fact_for63 (n : coq_type ml_int) : M (coq_type ml_int) := + do v <- cnew ml_int 1%int63; + do _ <- + (do u <- Ret 1%int63; + do v_1 <- Ret n; + forloop63 u v_1 + (fun i => + do v_1 <- (do v_1 <- cget v; Ret (mul v_1 i)); + cput v v_1)); + cget v. + +Definition int_of_nat n := Uint63.of_Z (Z.of_nat n). + +Lemma int_of_nat_succ : {morph int_of_nat : x / x.+1 >-> Uint63.succ x}. +Admitted. + +Lemma int_of_nat_mul : {morph int_of_nat : x y / x * y >-> mul x y}. +Admitted. + +Section fact_for63_ok. +Variable n : nat. +Hypothesis Hn : n < nat_of_uint Sint63.max_int. + +Let ltsb_succ : ltsb (int_of_nat n) (Uint63.succ (int_of_nat n)). +Admitted. +Lemma lesb_subr m : m < n -> lesb (int_of_nat (n - m)) (int_of_nat n). +Admitted. + +Theorem fact_for63_ok : + crun (fact_for63 (int_of_nat n)) = Some (int_of_nat (fact_rec n)). +Proof. +rewrite /fact_for63. +under eq_bind do rewrite !bindA !bindretf. +set fn := int_of_nat (fact_rec n). +transitivity (crun (cnew ml_int fn >> Ret fn : M _)); + last by rewrite crunret // crunnew0. +congr crun. +have {1}-> : (1 = int_of_nat 1)%int63 by []. +have -> : (1 = Uint63.succ (int_of_nat 0))%int63 by []. +rewrite -/(fact_rec 0). +pose m := n. +have -> : 0 = n - m by rewrite subnn. +have : m <= n by []. +elim: m => [|m IH] mn. + rewrite subn0 /forloop63 ltsb_succ. + under eq_bind do rewrite bindretf -cgetret. + by rewrite cnewget. +rewrite -int_of_nat_succ subnSK //. +under eq_bind do rewrite forloop63S !(lesb_subr,bindA) //. +rewrite cnewget. +under eq_bind do rewrite bindretf. +rewrite cnewput -IH (ltnW,subnS) // -int_of_nat_mul mulnC -(@prednK (n-m)) //. +by rewrite lt0n subn_eq0 -ltnNge. +Qed. +End fact_for63_ok. End fact_for_int63. From efd60c4382062ed4270e954f9c25810d2fcf6320 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 24 Feb 2023 19:19:49 +0900 Subject: [PATCH 28/82] use forloop630 --- example_typed_store.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 8e8ceca8..a5838eb3 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -357,8 +357,8 @@ pose m := n. have -> : 0 = n - m by rewrite subnn. have : m <= n by []. elim: m => [|m IH] mn. - rewrite subn0 /forloop63 ltsb_succ. - under eq_bind do rewrite bindretf -cgetret. + rewrite subn0. + under eq_bind do rewrite forloop630 (ltsb_succ,bindretf) // -cgetret. by rewrite cnewget. rewrite -int_of_nat_succ subnSK //. under eq_bind do rewrite forloop63S !(lesb_subr,bindA) //. From e398a03c76ba849a5442edc9e97999578bf3307c Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Sat, 25 Feb 2023 16:19:51 +0900 Subject: [PATCH 29/82] use conversion names from refproof --- example_typed_store.v | 34 +++++++++++++++++----------------- 1 file changed, 17 insertions(+), 17 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index a5838eb3..3fae3301 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -265,7 +265,7 @@ Section fact_for_int63. Variable M : typedStoreMonad. Notation coq_type := (@MLtypes63.coq_type M). -Definition nat_of_uint (n : int) : nat := +Definition uint2N (n : int) : nat := if Uint63.to_Z n is Zpos pos then Pos.to_nat pos else 0. Notation "'do' x <- m ; e" := (m >>= (fun x => e)) @@ -273,7 +273,7 @@ Notation "'do' x <- m ; e" := (m >>= (fun x => e)) Definition forloop63 (n_1 n_2 : int) (b : int -> M unit) : M unit := if Sint63.ltb n_2 n_1 then Ret tt else - iter (nat_of_uint (sub n_2 n_1)).+1 + iter (uint2N (sub n_2 n_1)).+1 (fun (m : M int) => do i <- m; do _ <- b i; Ret (Uint63.succ i)) (Ret n_1) >> Ret tt. @@ -291,8 +291,8 @@ Proof. by elim: n m2 => // n IH m2; rewrite iterS IH !bindA. Qed. Lemma lesb_ltsb_eq m n : lesb m n -> ltsb n (Uint63.succ m) -> m = n. Admitted. -Lemma nat_of_uint_sub_succ m n : lesb m n -> lesb (Uint63.succ m) n -> - nat_of_uint (sub n m) = (nat_of_uint (sub n (Uint63.succ m))).+1. +Lemma uint2N_sub_succ m n : lesb m n -> lesb (Uint63.succ m) n -> + uint2N (sub n m) = (uint2N (sub n (Uint63.succ m))).+1. Admitted. Lemma forloop63S m n (f : int -> M unit) : @@ -304,7 +304,7 @@ case: ifPn => m1n. rewrite (lesb_ltsb_eq _ _ mn m1n). by rewrite Sint63.sub_of_Z Z.sub_diag /= !(bindA,bindretf). rewrite ltsbNlesb negbK in m1n. -rewrite nat_of_uint_sub_succ //. +rewrite uint2N_sub_succ //. by rewrite iterSr bindretf !bindA iter_bind !bindA. Qed. @@ -323,35 +323,35 @@ Definition fact_for63 (n : coq_type ml_int) : M (coq_type ml_int) := cput v v_1)); cget v. -Definition int_of_nat n := Uint63.of_Z (Z.of_nat n). +Definition N2int n := Uint63.of_Z (Z.of_nat n). -Lemma int_of_nat_succ : {morph int_of_nat : x / x.+1 >-> Uint63.succ x}. +Lemma N2int_succ : {morph N2int : x / x.+1 >-> Uint63.succ x}. Admitted. -Lemma int_of_nat_mul : {morph int_of_nat : x y / x * y >-> mul x y}. +Lemma N2int_mul : {morph N2int : x y / x * y >-> mul x y}. Admitted. Section fact_for63_ok. Variable n : nat. -Hypothesis Hn : n < nat_of_uint Sint63.max_int. +Hypothesis Hn : n < uint2N Sint63.max_int. -Let ltsb_succ : ltsb (int_of_nat n) (Uint63.succ (int_of_nat n)). +Lemma ltsb_succ : ltsb (N2int n) (Uint63.succ (N2int n)). Admitted. -Lemma lesb_subr m : m < n -> lesb (int_of_nat (n - m)) (int_of_nat n). +Lemma lesb_subr m : m < n -> lesb (N2int (n - m)) (N2int n). Admitted. Theorem fact_for63_ok : - crun (fact_for63 (int_of_nat n)) = Some (int_of_nat (fact_rec n)). + crun (fact_for63 (N2int n)) = Some (N2int (fact_rec n)). Proof. rewrite /fact_for63. under eq_bind do rewrite !bindA !bindretf. -set fn := int_of_nat (fact_rec n). +set fn := N2int (fact_rec n). transitivity (crun (cnew ml_int fn >> Ret fn : M _)); last by rewrite crunret // crunnew0. congr crun. -have {1}-> : (1 = int_of_nat 1)%int63 by []. -have -> : (1 = Uint63.succ (int_of_nat 0))%int63 by []. +have {1}-> : (1 = N2int 1)%int63 by []. +have -> : (1 = Uint63.succ (N2int 0))%int63 by []. rewrite -/(fact_rec 0). pose m := n. have -> : 0 = n - m by rewrite subnn. @@ -360,11 +360,11 @@ elim: m => [|m IH] mn. rewrite subn0. under eq_bind do rewrite forloop630 (ltsb_succ,bindretf) // -cgetret. by rewrite cnewget. -rewrite -int_of_nat_succ subnSK //. +rewrite -N2int_succ subnSK //. under eq_bind do rewrite forloop63S !(lesb_subr,bindA) //. rewrite cnewget. under eq_bind do rewrite bindretf. -rewrite cnewput -IH (ltnW,subnS) // -int_of_nat_mul mulnC -(@prednK (n-m)) //. +rewrite cnewput -IH (ltnW,subnS) // -N2int_mul mulnC -(@prednK (n-m)) //. by rewrite lt0n subn_eq0 -ltnNge. Qed. End fact_for63_ok. From 376a3b0a17db1e50f6b7a51fbc23b6352bea3cb5 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Sat, 25 Feb 2023 16:49:23 +0900 Subject: [PATCH 30/82] reorder script for clarity --- example_typed_store.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/example_typed_store.v b/example_typed_store.v index 3fae3301..2c90f32f 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -351,8 +351,8 @@ transitivity (crun (cnew ml_int fn >> Ret fn : M _)); last by rewrite crunret // crunnew0. congr crun. have {1}-> : (1 = N2int 1)%int63 by []. -have -> : (1 = Uint63.succ (N2int 0))%int63 by []. rewrite -/(fact_rec 0). +have -> : (1 = Uint63.succ (N2int 0))%int63 by []. pose m := n. have -> : 0 = n - m by rewrite subnn. have : m <= n by []. From 0e04c6f83ae2ef9f931a7b9400bf9b9aec7ffa37 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 7 Mar 2023 00:30:46 +0900 Subject: [PATCH 31/82] fix 63-bit proof --- example_typed_store.v | 71 ++++++++++++++++++++++++++++++++++++------- 1 file changed, 60 insertions(+), 11 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 2c90f32f..4136d1db 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -283,31 +283,74 @@ case/boolP: (lesb n m) => /Sint63.lebP nm; apply/Sint63.ltbP => /=; by [apply Z.le_ngt | apply Z.nle_gt]. Qed. +Lemma ltsbNltsbS m n : ltsb m n -> ~ ltsb n (Uint63.succ m). +Proof. +move/Sint63.ltbP => mn. +move/Sint63.ltbP/Zlt_not_le; elim. +rewrite Sint63.succ_spec Sint63.cmod_small. + by apply/Zlt_le_succ. +split. + apply (Z.le_trans _ (Sint63.to_Z m)). by case: (Sint63.to_Z_bounded m). + by apply Z.le_succ_diag_r. +apply (Z.le_lt_trans _ (Sint63.to_Z n)). + by apply/Zlt_le_succ. +apply (Z.le_lt_trans _ (Sint63.to_Z Sint63.max_int)). + by case: (Sint63.to_Z_bounded n). +rewrite Sint63.to_Z_max. by apply Z.lt_sub_pos. +Qed. + +Lemma ltsbW m n : ltsb m n -> lesb m n. +Proof. move/Sint63.ltbP/Z.lt_le_incl => mn; by apply/Sint63.lebP. Qed. + Lemma iter_bind T n (f : T -> M T) (m1 : M unit) m2 : iter n (fun (m : M T) => m >>= f) (m1 >> m2) = m1 >> iter n (fun (m : M T) => m >>= f) m2. Proof. by elim: n m2 => // n IH m2; rewrite iterS IH !bindA. Qed. Lemma lesb_ltsb_eq m n : lesb m n -> ltsb n (Uint63.succ m) -> m = n. -Admitted. +Proof. +move/Sint63.lebP => mn /Sint63.ltbP nSm. +move: (nSm). +rewrite Sint63.succ_of_Z -Sint63.is_int; last first. + split. + apply Z.le_le_succ_r. + by case: (Sint63.to_Z_bounded m). + apply Zlt_le_succ, (Z.le_lt_trans _ _ _ mn), (Z.lt_le_trans _ _ _ nSm). + by case: (Sint63.to_Z_bounded (Uint63.succ m)). +move/Zlt_le_succ/Zsucc_le_reg => nm. +by apply Sint63.to_Z_inj, Zle_antisym. +Qed. -Lemma uint2N_sub_succ m n : lesb m n -> lesb (Uint63.succ m) n -> +Lemma uint2N_sub_succ m n : ltsb m n -> uint2N (sub n m) = (uint2N (sub n (Uint63.succ m))).+1. +Proof. +rewrite /uint2N !Uint63.sub_spec Uint63.succ_spec. +rewrite [in RHS]Zminus_mod Zmod_mod -Zminus_mod. +rewrite Z.sub_add_distr. +move/Sint63.ltbP => mn. Admitted. Lemma forloop63S m n (f : int -> M unit) : - lesb m n -> forloop63 m n f = f m >> forloop63 (Uint63.succ m) n f. + ltsb m n -> forloop63 m n f = f m >> forloop63 (Uint63.succ m) n f. Proof. rewrite /forloop63 => mn. -rewrite ltsbNlesb mn /=. -case: ifPn => m1n. - rewrite (lesb_ltsb_eq _ _ mn m1n). - by rewrite Sint63.sub_of_Z Z.sub_diag /= !(bindA,bindretf). -rewrite ltsbNlesb negbK in m1n. +rewrite ltsbNlesb (ltsbW _ _ mn) /=. +case: ifPn => nSm. + by elim (ltsbNltsbS _ _ mn). +rewrite ltsbNlesb negbK in nSm. rewrite uint2N_sub_succ //. by rewrite iterSr bindretf !bindA iter_bind !bindA. Qed. +Lemma forloop631 m (f : int -> M unit) : + forloop63 m m f = f m. +Proof. rewrite /forloop63. +case: (Sint63.ltbP m m) => [/Z.lt_irrefl // | _]. +rewrite /= bindA. +rewrite /uint2N Uint63.sub_spec Z.sub_diag Zmod_0_l /=. +by rewrite !(bindretf,bindA) bindmskip. +Qed. + Lemma forloop630 m n (f : int -> M unit) : ltsb n m -> forloop63 m n f = skip. Proof. by rewrite /forloop63 => ->. Qed. @@ -338,7 +381,7 @@ Hypothesis Hn : n < uint2N Sint63.max_int. Lemma ltsb_succ : ltsb (N2int n) (Uint63.succ (N2int n)). Admitted. -Lemma lesb_subr m : m < n -> lesb (N2int (n - m)) (N2int n). +Lemma ltsb_subr m : m.+1 < n -> ltsb (N2int (n - m.+1)) (N2int n). Admitted. Theorem fact_for63_ok : @@ -361,10 +404,16 @@ elim: m => [|m IH] mn. under eq_bind do rewrite forloop630 (ltsb_succ,bindretf) // -cgetret. by rewrite cnewget. rewrite -N2int_succ subnSK //. -under eq_bind do rewrite forloop63S !(lesb_subr,bindA) //. +case: m IH mn => [|m] IH mn. + under eq_bind do rewrite subn0 forloop631 !(ltsb_subr,bindA) //. + rewrite cnewget. + under eq_bind do rewrite bindretf -cgetret. + rewrite cnewput -N2int_mul mulnC -{1}(@prednK n) // cnewget subn1. + by rewrite -/(fact_rec n.-1.+1) prednK. +under eq_bind do rewrite forloop63S !(ltsb_subr,bindA) //. rewrite cnewget. under eq_bind do rewrite bindretf. -rewrite cnewput -IH (ltnW,subnS) // -N2int_mul mulnC -(@prednK (n-m)) //. +rewrite cnewput -IH (ltnW,subnS) // -N2int_mul mulnC -(@prednK (n-m.+1)) //. by rewrite lt0n subn_eq0 -ltnNge. Qed. End fact_for63_ok. From a91b86546b1f4074d8c4ea4804020e9bc3b263e9 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 7 Mar 2023 12:41:58 +0900 Subject: [PATCH 32/82] down to one admit --- example_typed_store.v | 80 ++++++++++++++++++++++++++++++++++++++----- 1 file changed, 72 insertions(+), 8 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 4136d1db..d49f24b0 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -321,13 +321,38 @@ move/Zlt_le_succ/Zsucc_le_reg => nm. by apply Sint63.to_Z_inj, Zle_antisym. Qed. +Lemma uint63_min n : (0 <= Uint63.to_Z n)%Z. +Proof. by case: (Uint63.to_Z_bounded n). Qed. + +Lemma uint63_max n : (Uint63.to_Z n < Uint63.wB)%Z. +Proof. by case: (Uint63.to_Z_bounded n). Qed. + Lemma uint2N_sub_succ m n : ltsb m n -> uint2N (sub n m) = (uint2N (sub n (Uint63.succ m))).+1. Proof. rewrite /uint2N !Uint63.sub_spec Uint63.succ_spec. -rewrite [in RHS]Zminus_mod Zmod_mod -Zminus_mod. -rewrite Z.sub_add_distr. -move/Sint63.ltbP => mn. +rewrite [in RHS]Zminus_mod Zmod_mod -Zminus_mod Z.sub_add_distr. +move/Sint63.ltbP. +rewrite /Sint63.to_Z. +case: ifP => /Uint63.ltbP Hm; case: ifP => /Uint63.ltbP Hn. +- move=> mn. + rewrite !Zmod_small; first last. + + split. by apply Zle_minus_le_0, Z.lt_le_incl. + rewrite -(subZ0 Uint63.wB). + by apply /Z.sub_lt_le_mono /uint63_min /uint63_max. + + split. by apply/leZsub1/Z.lt_0_sub. + rewrite -Z.sub_add_distr -(subZ0 Uint63.wB). + apply Z.sub_lt_le_mono. by apply uint63_max. + apply Z.add_nonneg_nonneg => //. by apply uint63_min. + case Hnm1: (Uint63.to_Z n - Uint63.to_Z m - 1)%Z. + + by move/Zminus_eq: Hnm1 => ->. + + move/Z.sub_move_r: Hnm1 => -> /=. + by rewrite -Pos2Nat.inj_succ Pplus_one_succ_r. + + move/Z.lt_0_sub: mn. + move/Z.sub_move_r: Hnm1 => -> /ltZadd1 /Zle_not_lt; elim. + by apply Pos2Z.neg_is_neg. +- move/Zlt_not_le; elim. + by apply /(Z.le_trans _ 0) /uint63_min /Z.opp_nonpos_nonneg /uint63_min. Admitted. Lemma forloop63S m n (f : int -> M unit) : @@ -369,20 +394,59 @@ Definition fact_for63 (n : coq_type ml_int) : M (coq_type ml_int) := Definition N2int n := Uint63.of_Z (Z.of_nat n). Lemma N2int_succ : {morph N2int : x / x.+1 >-> Uint63.succ x}. -Admitted. +Proof. +move=> x; apply Uint63.to_Z_inj; rewrite Uint63.succ_spec !Uint63.of_Z_spec. +by rewrite Zplus_mod /= Zpos_P_of_succ_nat /Z.succ Zplus_mod Zmod_mod. +Qed. Lemma N2int_mul : {morph N2int : x y / x * y >-> mul x y}. -Admitted. +Proof. +move=> x y; apply Uint63.to_Z_inj. +by rewrite Uint63.mul_spec !Uint63.of_Z_spec Nat2Z.inj_mul Zmult_mod. +Qed. + +Lemma N2int_bounded n : + (Z.of_nat n <= Sint63.to_Z Sint63.max_int)%Z -> + (Sint63.to_Z Sint63.min_int <= Z.of_nat n <= Sint63.to_Z Sint63.max_int)%Z. +Proof. +split => //. +apply (Z.le_trans _ 0). + rewrite -[0%Z]/(Sint63.to_Z 0). + by case: (Sint63.to_Z_bounded 0). +by apply Zle_0_nat. +Qed. Section fact_for63_ok. Variable n : nat. -Hypothesis Hn : n < uint2N Sint63.max_int. +Hypothesis Hn : (Z.of_nat n < Sint63.to_Z Sint63.max_int)%Z. + +Let n_bounded : + (Sint63.to_Z Sint63.min_int <= Z.of_nat n <= Sint63.to_Z Sint63.max_int)%Z. +Proof. by apply N2int_bounded, Z.lt_le_incl. Qed. Lemma ltsb_succ : ltsb (N2int n) (Uint63.succ (N2int n)). -Admitted. +Proof. +apply/Sint63.ltbP. +rewrite Sint63.succ_spec Sint63.cmod_small. + by apply/Zle_lt_succ/Z.le_refl. +split. + apply leZ_addr => //; by case: (Sint63.to_Z_bounded (N2int n)). +apply Z.lt_add_lt_sub_r; by rewrite -Sint63.is_int. +Qed. Lemma ltsb_subr m : m.+1 < n -> ltsb (N2int (n - m.+1)) (N2int n). -Admitted. +Proof. +move=> Smn. +apply/Sint63.ltbP. +have Hm : n - m.+1 < n. + rewrite ltn_subLR. + by rewrite addSn ltnS leq_addl. + by apply ltnW. +rewrite /N2int -!Sint63.is_int //. +- by apply/inj_lt/ltP. +- move/ltP/inj_lt in Hm. + by split; apply N2int_bounded, Z.lt_le_incl, (Z.lt_trans _ _ _ Hm). +Qed. Theorem fact_for63_ok : crun (fact_for63 (N2int n)) = Some (N2int (fact_rec n)). From 079ce443d154323e9955422f5c588979c75fc90b Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 7 Mar 2023 17:29:00 +0900 Subject: [PATCH 33/82] finish proof of uint2N_sub_succ --- example_typed_store.v | 77 +++++++++++++++++++++++++++++-------------- 1 file changed, 52 insertions(+), 25 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index d49f24b0..a02660dc 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -327,33 +327,60 @@ Proof. by case: (Uint63.to_Z_bounded n). Qed. Lemma uint63_max n : (Uint63.to_Z n < Uint63.wB)%Z. Proof. by case: (Uint63.to_Z_bounded n). Qed. +Lemma succ_pred m n : sub n (Uint63.succ m) = Uint63.pred (sub n m). +Proof. +apply Uint63.to_Z_inj. +rewrite Uint63.sub_spec Uint63.succ_spec Uint63.pred_spec Uint63.sub_spec. +rewrite Zminus_mod Zmod_mod -Zminus_mod Z.sub_add_distr. +apply/esym. +by rewrite Zminus_mod Zmod_mod -Zminus_mod. +Qed. + +Lemma uint2N_pred n : n <> 0%int63 -> uint2N n = (uint2N (Uint63.pred n)).+1. +Proof. +move=> Hn. +rewrite /uint2N Uint63.pred_spec. +case HnZ: (Uint63.to_Z n) => [|m|m]. +- rewrite (_ : 0 = Uint63.to_Z 0)%Z // in HnZ. + move/Uint63.to_Z_inj in HnZ. + by elim Hn. +- have Hm1 : (0 <= Z.pos m - 1 < Uint63.wB)%Z. + split. by apply leZsub1, Pos2Z.is_pos. + apply (Z.lt_trans _ (Z.pos m)). + by apply leZsub1, Z.le_refl. + rewrite -HnZ; by apply uint63_max. + rewrite Zmod_small //. + case HmZ: (Z.pos m - 1)%Z => [|p|p]. + + by move/Z.sub_move_r: HmZ => /= [] ->. + + apply Nat2Z.inj => /=. + rewrite positive_nat_Z Pos2SuccNat.id_succ Pos2Z.inj_succ -HmZ. + by rewrite (Z.succ_pred (Z.pos m)). + + by destruct m. +- move: (uint63_min n). + rewrite HnZ => /Zle_not_lt; elim. + by apply Zlt_neg_0. +Qed. + Lemma uint2N_sub_succ m n : ltsb m n -> uint2N (sub n m) = (uint2N (sub n (Uint63.succ m))).+1. Proof. -rewrite /uint2N !Uint63.sub_spec Uint63.succ_spec. -rewrite [in RHS]Zminus_mod Zmod_mod -Zminus_mod Z.sub_add_distr. -move/Sint63.ltbP. -rewrite /Sint63.to_Z. -case: ifP => /Uint63.ltbP Hm; case: ifP => /Uint63.ltbP Hn. -- move=> mn. - rewrite !Zmod_small; first last. - + split. by apply Zle_minus_le_0, Z.lt_le_incl. - rewrite -(subZ0 Uint63.wB). - by apply /Z.sub_lt_le_mono /uint63_min /uint63_max. - + split. by apply/leZsub1/Z.lt_0_sub. - rewrite -Z.sub_add_distr -(subZ0 Uint63.wB). - apply Z.sub_lt_le_mono. by apply uint63_max. - apply Z.add_nonneg_nonneg => //. by apply uint63_min. - case Hnm1: (Uint63.to_Z n - Uint63.to_Z m - 1)%Z. - + by move/Zminus_eq: Hnm1 => ->. - + move/Z.sub_move_r: Hnm1 => -> /=. - by rewrite -Pos2Nat.inj_succ Pplus_one_succ_r. - + move/Z.lt_0_sub: mn. - move/Z.sub_move_r: Hnm1 => -> /ltZadd1 /Zle_not_lt; elim. - by apply Pos2Z.neg_is_neg. -- move/Zlt_not_le; elim. - by apply /(Z.le_trans _ 0) /uint63_min /Z.opp_nonpos_nonneg /uint63_min. -Admitted. +move=> mn. +rewrite succ_pred uint2N_pred //. +rewrite Sint63.sub_of_Z => /(f_equal Uint63.to_Z). +rewrite Uint63.of_Z_spec. +move/Sint63.ltbP in mn. +rewrite Zmod_small. + rewrite Z.sub_move_r /= => nm. + rewrite nm in mn. + by move/Z.lt_irrefl in mn. +split. by apply Zle_minus_le_0, Z.lt_le_incl. +apply + (Z.le_lt_trans _ (Sint63.to_Z Sint63.max_int - Sint63.to_Z Sint63.min_int))%Z. + apply leZ_sub. + by case: (Sint63.to_Z_bounded n). + by case: (Sint63.to_Z_bounded m). +done. +Qed. Lemma forloop63S m n (f : int -> M unit) : ltsb m n -> forloop63 m n f = f m >> forloop63 (Uint63.succ m) n f. @@ -472,7 +499,7 @@ case: m IH mn => [|m] IH mn. under eq_bind do rewrite subn0 forloop631 !(ltsb_subr,bindA) //. rewrite cnewget. under eq_bind do rewrite bindretf -cgetret. - rewrite cnewput -N2int_mul mulnC -{1}(@prednK n) // cnewget subn1. + rewrite cnewput -N2int_mul mulnC -{1}(prednK mn) cnewget subn1. by rewrite -/(fact_rec n.-1.+1) prednK. under eq_bind do rewrite forloop63S !(ltsb_subr,bindA) //. rewrite cnewget. From 9beb5e48e73dd3077abb98c02824b46cd738312f Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 7 Mar 2023 18:02:58 +0900 Subject: [PATCH 34/82] extract lemmas --- example_typed_store.v | 31 ++++++++++++++++++++----------- 1 file changed, 20 insertions(+), 11 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index a02660dc..20567176 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -359,13 +359,24 @@ case HnZ: (Uint63.to_Z n) => [|m|m]. - move: (uint63_min n). rewrite HnZ => /Zle_not_lt; elim. by apply Zlt_neg_0. -Qed. +Qed. -Lemma uint2N_sub_succ m n : ltsb m n -> - uint2N (sub n m) = (uint2N (sub n (Uint63.succ m))).+1. +Lemma lesb_sub_bounded m n : + lesb m n -> (0 <= Sint63.to_Z n - Sint63.to_Z m < Uint63.wB)%Z. +Proof. +move/Sint63.lebP => mn. +split. by apply Zle_minus_le_0. +apply + (Z.le_lt_trans _ (Sint63.to_Z Sint63.max_int - Sint63.to_Z Sint63.min_int))%Z. + apply leZ_sub. + by case: (Sint63.to_Z_bounded n). + by case: (Sint63.to_Z_bounded m). +done. +Qed. + +Lemma ltsb_sub_neq0 m n : ltsb m n -> sub n m <> 0%int63. Proof. move=> mn. -rewrite succ_pred uint2N_pred //. rewrite Sint63.sub_of_Z => /(f_equal Uint63.to_Z). rewrite Uint63.of_Z_spec. move/Sint63.ltbP in mn. @@ -373,15 +384,13 @@ rewrite Zmod_small. rewrite Z.sub_move_r /= => nm. rewrite nm in mn. by move/Z.lt_irrefl in mn. -split. by apply Zle_minus_le_0, Z.lt_le_incl. -apply - (Z.le_lt_trans _ (Sint63.to_Z Sint63.max_int - Sint63.to_Z Sint63.min_int))%Z. - apply leZ_sub. - by case: (Sint63.to_Z_bounded n). - by case: (Sint63.to_Z_bounded m). -done. +by apply /lesb_sub_bounded /Sint63.lebP /Z.lt_le_incl. Qed. +Lemma uint2N_sub_succ m n : ltsb m n -> + uint2N (sub n m) = (uint2N (sub n (Uint63.succ m))).+1. +Proof. move/ltsb_sub_neq0 => mn. by rewrite succ_pred uint2N_pred. Qed. + Lemma forloop63S m n (f : int -> M unit) : ltsb m n -> forloop63 m n f = f m >> forloop63 (Uint63.succ m) n f. Proof. From 4a290fd4e39ab809693009fdc49b6869831b4258 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Wed, 8 Mar 2023 09:37:55 +0900 Subject: [PATCH 35/82] separate int63-related lemmas --- example_typed_store.v | 167 +++++++++++++++++++++--------------------- 1 file changed, 85 insertions(+), 82 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 20567176..6bc3c836 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -232,50 +232,10 @@ End fibonacci. Require Import PrimInt63. Require Sint63. -Module MLtypes63. -Local Definition ml_type_eq_dec := ml_type_eq_dec. -Local Definition ml_type := ml_type. -Local Definition undef := Undef. -Local Definition loc := loc. -Local Definition locT := [eqType of nat]. -Local Definition loc_id := @loc_id. - -Section with_monad. -Context [M : Type -> Type]. - -(* Generated type translation function *) -Fixpoint coq_type (T : ml_type) : Type := - match T with - | ml_int => int - | ml_bool => bool - | ml_unit => unit - | ml_arrow T1 T2 => coq_type T1 -> M (coq_type T2) - | ml_ref T1 => loc T1 - | ml_undef => undef_t - end. -End with_monad. -Local Definition ml_undef := ml_undef. -End MLtypes63. - -Module IMonadTS63 := hierarchy.MonadTypedStore (MLtypes63). -Import MLtypes63. -Import IMonadTS63. - -Section fact_for_int63. -Variable M : typedStoreMonad. -Notation coq_type := (@MLtypes63.coq_type M). - +Section Int63. Definition uint2N (n : int) : nat := if Uint63.to_Z n is Zpos pos then Pos.to_nat pos else 0. - -Notation "'do' x <- m ; e" := (m >>= (fun x => e)) - (at level 60, x name, m at level 200, e at level 60). - -Definition forloop63 (n_1 n_2 : int) (b : int -> M unit) : M unit := - if Sint63.ltb n_2 n_1 then Ret tt else - iter (uint2N (sub n_2 n_1)).+1 - (fun (m : M int) => do i <- m; do _ <- b i; Ret (Uint63.succ i)) - (Ret n_1) >> Ret tt. +Definition N2int n := Uint63.of_Z (Z.of_nat n). Lemma ltsbNlesb m n : ltsb m n = ~~ lesb n m. Proof. @@ -302,11 +262,6 @@ Qed. Lemma ltsbW m n : ltsb m n -> lesb m n. Proof. move/Sint63.ltbP/Z.lt_le_incl => mn; by apply/Sint63.lebP. Qed. -Lemma iter_bind T n (f : T -> M T) (m1 : M unit) m2 : - iter n (fun (m : M T) => m >>= f) (m1 >> m2) = - m1 >> iter n (fun (m : M T) => m >>= f) m2. -Proof. by elim: n m2 => // n IH m2; rewrite iterS IH !bindA. Qed. - Lemma lesb_ltsb_eq m n : lesb m n -> ltsb n (Uint63.succ m) -> m = n. Proof. move/Sint63.lebP => mn /Sint63.ltbP nSm. @@ -327,15 +282,6 @@ Proof. by case: (Uint63.to_Z_bounded n). Qed. Lemma uint63_max n : (Uint63.to_Z n < Uint63.wB)%Z. Proof. by case: (Uint63.to_Z_bounded n). Qed. -Lemma succ_pred m n : sub n (Uint63.succ m) = Uint63.pred (sub n m). -Proof. -apply Uint63.to_Z_inj. -rewrite Uint63.sub_spec Uint63.succ_spec Uint63.pred_spec Uint63.sub_spec. -rewrite Zminus_mod Zmod_mod -Zminus_mod Z.sub_add_distr. -apply/esym. -by rewrite Zminus_mod Zmod_mod -Zminus_mod. -Qed. - Lemma uint2N_pred n : n <> 0%int63 -> uint2N n = (uint2N (Uint63.pred n)).+1. Proof. move=> Hn. @@ -387,9 +333,90 @@ rewrite Zmod_small. by apply /lesb_sub_bounded /Sint63.lebP /Z.lt_le_incl. Qed. +Lemma sub_succ_pred m n : sub n (Uint63.succ m) = Uint63.pred (sub n m). +Proof. +apply Uint63.to_Z_inj. +rewrite Uint63.sub_spec Uint63.succ_spec Uint63.pred_spec Uint63.sub_spec. +rewrite Zminus_mod Zmod_mod -Zminus_mod Z.sub_add_distr. +apply/esym. +by rewrite Zminus_mod Zmod_mod -Zminus_mod. +Qed. + Lemma uint2N_sub_succ m n : ltsb m n -> uint2N (sub n m) = (uint2N (sub n (Uint63.succ m))).+1. -Proof. move/ltsb_sub_neq0 => mn. by rewrite succ_pred uint2N_pred. Qed. +Proof. move/ltsb_sub_neq0 => mn. by rewrite sub_succ_pred uint2N_pred. Qed. + +Lemma N2int_succ : {morph N2int : x / x.+1 >-> Uint63.succ x}. +Proof. +move=> x; apply Uint63.to_Z_inj; rewrite Uint63.succ_spec !Uint63.of_Z_spec. +by rewrite Zplus_mod /= Zpos_P_of_succ_nat /Z.succ Zplus_mod Zmod_mod. +Qed. + +Lemma N2int_mul : {morph N2int : x y / x * y >-> mul x y}. +Proof. +move=> x y; apply Uint63.to_Z_inj. +by rewrite Uint63.mul_spec !Uint63.of_Z_spec Nat2Z.inj_mul Zmult_mod. +Qed. + +Lemma N2int_bounded n : + (Z.of_nat n <= Sint63.to_Z Sint63.max_int)%Z -> + (Sint63.to_Z Sint63.min_int <= Z.of_nat n <= Sint63.to_Z Sint63.max_int)%Z. +Proof. +split => //. +apply (Z.le_trans _ 0). + rewrite -[0%Z]/(Sint63.to_Z 0). + by case: (Sint63.to_Z_bounded 0). +by apply Zle_0_nat. +Qed. +End Int63. + +Module MLtypes63. +Local Definition ml_type_eq_dec := ml_type_eq_dec. +Local Definition ml_type := ml_type. +Local Definition undef := Undef. +Local Definition loc := loc. +Local Definition locT := [eqType of nat]. +Local Definition loc_id := @loc_id. + +Section with_monad. +Context [M : Type -> Type]. + +(* Generated type translation function *) +Fixpoint coq_type (T : ml_type) : Type := + match T with + | ml_int => int + | ml_bool => bool + | ml_unit => unit + | ml_arrow T1 T2 => coq_type T1 -> M (coq_type T2) + | ml_ref T1 => loc T1 + | ml_undef => undef_t + end. +End with_monad. +Local Definition ml_undef := ml_undef. +End MLtypes63. + +Module IMonadTS63 := hierarchy.MonadTypedStore (MLtypes63). +Import MLtypes63. +Import IMonadTS63. + +Section fact_for_int63. +Variable M : typedStoreMonad. +Notation coq_type := (@MLtypes63.coq_type M). + +Notation "'do' x <- m ; e" := (m >>= (fun x => e)) + (at level 60, x name, m at level 200, e at level 60). + +Section forloop63. +Definition forloop63 (n_1 n_2 : int) (b : int -> M unit) : M unit := + if Sint63.ltb n_2 n_1 then Ret tt else + iter (uint2N (sub n_2 n_1)).+1 + (fun (m : M int) => do i <- m; do _ <- b i; Ret (Uint63.succ i)) + (Ret n_1) >> Ret tt. + +Lemma iter_bind T n (f : T -> M T) (m1 : M unit) m2 : + iter n (fun (m : M T) => m >>= f) (m1 >> m2) = + m1 >> iter n (fun (m : M T) => m >>= f) m2. +Proof. by elim: n m2 => // n IH m2; rewrite iterS IH !bindA. Qed. Lemma forloop63S m n (f : int -> M unit) : ltsb m n -> forloop63 m n f = f m >> forloop63 (Uint63.succ m) n f. @@ -415,6 +442,7 @@ Qed. Lemma forloop630 m n (f : int -> M unit) : ltsb n m -> forloop63 m n f = skip. Proof. by rewrite /forloop63 => ->. Qed. +End forloop63. Definition fact_for63 (n : coq_type ml_int) : M (coq_type ml_int) := do v <- cnew ml_int 1%int63; @@ -427,31 +455,6 @@ Definition fact_for63 (n : coq_type ml_int) : M (coq_type ml_int) := cput v v_1)); cget v. -Definition N2int n := Uint63.of_Z (Z.of_nat n). - -Lemma N2int_succ : {morph N2int : x / x.+1 >-> Uint63.succ x}. -Proof. -move=> x; apply Uint63.to_Z_inj; rewrite Uint63.succ_spec !Uint63.of_Z_spec. -by rewrite Zplus_mod /= Zpos_P_of_succ_nat /Z.succ Zplus_mod Zmod_mod. -Qed. - -Lemma N2int_mul : {morph N2int : x y / x * y >-> mul x y}. -Proof. -move=> x y; apply Uint63.to_Z_inj. -by rewrite Uint63.mul_spec !Uint63.of_Z_spec Nat2Z.inj_mul Zmult_mod. -Qed. - -Lemma N2int_bounded n : - (Z.of_nat n <= Sint63.to_Z Sint63.max_int)%Z -> - (Sint63.to_Z Sint63.min_int <= Z.of_nat n <= Sint63.to_Z Sint63.max_int)%Z. -Proof. -split => //. -apply (Z.le_trans _ 0). - rewrite -[0%Z]/(Sint63.to_Z 0). - by case: (Sint63.to_Z_bounded 0). -by apply Zle_0_nat. -Qed. - Section fact_for63_ok. Variable n : nat. Hypothesis Hn : (Z.of_nat n < Sint63.to_Z Sint63.max_int)%Z. From 0817a4fe95f5e5c6301c64e8dfd88715fc257ae2 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Wed, 8 Mar 2023 15:01:49 +0900 Subject: [PATCH 36/82] factorize --- example_typed_store.v | 27 +++++++++------------------ 1 file changed, 9 insertions(+), 18 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 6bc3c836..ff728703 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -243,26 +243,10 @@ case/boolP: (lesb n m) => /Sint63.lebP nm; apply/Sint63.ltbP => /=; by [apply Z.le_ngt | apply Z.nle_gt]. Qed. -Lemma ltsbNltsbS m n : ltsb m n -> ~ ltsb n (Uint63.succ m). -Proof. -move/Sint63.ltbP => mn. -move/Sint63.ltbP/Zlt_not_le; elim. -rewrite Sint63.succ_spec Sint63.cmod_small. - by apply/Zlt_le_succ. -split. - apply (Z.le_trans _ (Sint63.to_Z m)). by case: (Sint63.to_Z_bounded m). - by apply Z.le_succ_diag_r. -apply (Z.le_lt_trans _ (Sint63.to_Z n)). - by apply/Zlt_le_succ. -apply (Z.le_lt_trans _ (Sint63.to_Z Sint63.max_int)). - by case: (Sint63.to_Z_bounded n). -rewrite Sint63.to_Z_max. by apply Z.lt_sub_pos. -Qed. - Lemma ltsbW m n : ltsb m n -> lesb m n. Proof. move/Sint63.ltbP/Z.lt_le_incl => mn; by apply/Sint63.lebP. Qed. -Lemma lesb_ltsb_eq m n : lesb m n -> ltsb n (Uint63.succ m) -> m = n. +Lemma lesb_ltsbS_eq m n : lesb m n -> ltsb n (Uint63.succ m) -> m = n. Proof. move/Sint63.lebP => mn /Sint63.ltbP nSm. move: (nSm). @@ -276,6 +260,13 @@ move/Zlt_le_succ/Zsucc_le_reg => nm. by apply Sint63.to_Z_inj, Zle_antisym. Qed. +(* The opposite is not true ! (n = max_int) *) +Lemma ltsbS_lesb m n : ltsb m (Uint63.succ n) -> lesb m n. +Proof. +rewrite -[lesb _ _]negbK -ltsbNlesb => nSm; apply/negP => /[dup] /ltsbW nm. +by rewrite (lesb_ltsbS_eq n m) // => /Sint63.ltbP /Z.lt_irrefl. +Qed. + Lemma uint63_min n : (0 <= Uint63.to_Z n)%Z. Proof. by case: (Uint63.to_Z_bounded n). Qed. @@ -424,7 +415,7 @@ Proof. rewrite /forloop63 => mn. rewrite ltsbNlesb (ltsbW _ _ mn) /=. case: ifPn => nSm. - by elim (ltsbNltsbS _ _ mn). + by move: mn; rewrite ltsbNlesb => /negP; elim; apply ltsbS_lesb. rewrite ltsbNlesb negbK in nSm. rewrite uint2N_sub_succ //. by rewrite iterSr bindretf !bindA iter_bind !bindA. From 9cd54eca6610cad520bbd54f88778fa114fe0dfb Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 23 May 2023 16:22:54 +0900 Subject: [PATCH 37/82] HB.pack test --- typed_store_model.v | 30 ++++++++++++++++++++---------- 1 file changed, 20 insertions(+), 10 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 153d60dc..52c00395 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -118,18 +118,29 @@ Local Notation "m >> f" := (bind m (fun=> f)).*) Definition actm (A B : UU0) (f : A -> B) (m : M A) : M B := mkActo (actm _ _ f (ofActo m)). -Definition isFunctorTS : isFunctor.axioms_ M. -exists actm. -- move=> A. apply boolp.funext => -[] {A} A m. - by rewrite /actm functor_id. -- move=> A B C g h. apply boolp.funext => m. - case: m h => {A} A m h /=. - by rewrite /actm /= functor_o. +Let actm_id : FunctorLaws.id actm. +Proof. +move=> A. apply boolp.funext => -[] {A} A m. +by rewrite /actm functor_id. +Qed. + +Let actm_comp : FunctorLaws.comp actm. +Proof. +move=> A B C g h. apply boolp.funext => m. +case: m h => {A} A m h /=. +by rewrite /actm /= functor_o. Defined. +Lemma isFunctorTS : isFunctor.axioms_ M. +Proof. by exists actm; [exact: actm_id|exact: actm_comp]. Qed. (* NB: not Qed, o.w. ret_naturality fails *) + Unset Universe Checking. -Canonical Structure FunctorTS : Functor M := Functor.Class isFunctorTS. -Canonical Structure functorTS : functor := Functor.Pack isFunctorTS. +(*HB.instance Definition _ := isFunctor.Build acto actm_id actm_comp. fails*) +Definition tmp := isFunctor.Build acto actm_id actm_comp. +Definition functorTS : functor := HB.pack acto tmp. +Canonical Structure functorTS. +(*Canonical Structure FunctorTS : Functor M := Functor.Class isFunctorTS. +Canonical Structure functorTS : functor := Functor.Pack isFunctorTS.*) Set Universe Checking. Definition ret_naturality : naturality [the functor of idfun] M ret. @@ -140,7 +151,6 @@ Canonical Structure naturalityTSret := isNatural.Build _ _ _ ret_naturality. Definition ret' : [the functor of idfun] ~> [the functor of M] := @Nattrans.Pack _ _ ret (Nattrans.Class naturalityTSret). -Unset Universe Checking. Definition join' : [the functor of M \o M] ~~> [the functor of M] := fun _ m => bind m idfun. Set Universe Checking. From 7339390b4883bb07c065e76501c3cdeccea0eb04 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 23 May 2023 16:32:14 +0900 Subject: [PATCH 38/82] fix --- typed_store_model.v | 1 + 1 file changed, 1 insertion(+) diff --git a/typed_store_model.v b/typed_store_model.v index 52c00395..d92bf429 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -151,6 +151,7 @@ Canonical Structure naturalityTSret := isNatural.Build _ _ _ ret_naturality. Definition ret' : [the functor of idfun] ~> [the functor of M] := @Nattrans.Pack _ _ ret (Nattrans.Class naturalityTSret). +Unset Universe Checking. Definition join' : [the functor of M \o M] ~~> [the functor of M] := fun _ m => bind m idfun. Set Universe Checking. From e8f163380bdbf58f0e659bea78167397443e61e1 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 23 May 2023 16:32:07 +0900 Subject: [PATCH 39/82] wip --- example_typed_store.v | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/example_typed_store.v b/example_typed_store.v index ff728703..7201483e 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -311,6 +311,23 @@ apply done. Qed. +Lemma ltsb_neq m n : ltsb m n -> m <> n. +Proof. by move/Sint63.ltbP/[swap]/(f_equal Sint63.to_Z)-> =>/Z.lt_irrefl. Qed. + +(* +Lemma sub0_eq m n : sub m n = 0%int63 -> m = n. +Proof. +rewrite Sint63.sub_of_Z => /(f_equal Uint63.to_Z). +rewrite Uint63.of_Z_spec. +move/Sint63.ltbP in mn. +rewrite Zmod_small. + rewrite Z.sub_move_r /= => nm. + rewrite nm in mn. + by move/Z.lt_irrefl in mn. +by apply /lesb_sub_bounded /Sint63.lebP /Z.lt_le_incl. +Qed. +*) + Lemma ltsb_sub_neq0 m n : ltsb m n -> sub n m <> 0%int63. Proof. move=> mn. From 45f5e5952b88d712d8d179011bb8d74fc97ae9a4 Mon Sep 17 00:00:00 2001 From: Takafumi Saikawa Date: Tue, 23 May 2023 18:07:10 +0900 Subject: [PATCH 40/82] faster actm_bind --- typed_store_model.v | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index d92bf429..ea50ad8a 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -159,10 +159,8 @@ Set Universe Checking. Definition actm_bind (a b c : UU0) (f : a -> b) m (g : c -> M a) : (actm f) (bind m g) = bind m (actm f \o g). Proof. -congr mkActo. -case: m g => {}c m g /=. -rewrite bindE fmapE 2!bindA /=. -reflexivity. +rewrite /actm fmapE bindA. +by cbv. Qed. Lemma mkActoK A (m : MS _ _ A) : ofActo (mkActo m) = m. From e4df44b66c6dce757faeff0e1892e2ba2e08bdf3 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 23 May 2023 18:21:16 +0900 Subject: [PATCH 41/82] try exception --- monad_model.v | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/monad_model.v b/monad_model.v index 29921cd5..e307bab5 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1722,6 +1722,7 @@ Variant loc : ml_type -> Set := mkloc T : nat -> loc T. Parameter coq_type0 : ml_type -> Type. Parameter ml_undef : ml_type. Parameter undef : coq_type0 ml_undef. +(*Parameter ml_exn : ml_type.*) End MLTYweak. Module ModelTypedStore (MLtypes : MLTYweak). @@ -1744,8 +1745,9 @@ Record binding (M : Type -> Type) := Arguments mkbind {M}. Section typed_store. -Definition acto (T : UU0) : UU0 := - MS (seq (binding idfun)) [the monad of option_monad] T. +Definition acto : UU0 -> UU0 := + (*MX (coq_type0 ml_exn)*) + (MS (seq (binding idfun)) [the monad of option_monad]). Local Notation M := acto. (*Local Notation coq_type := (coq_type M).*) Local Notation coq_type' := (coq_type idfun). @@ -1800,8 +1802,7 @@ Let cnewget T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A) : Proof. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. +rewrite bindE /= bindE /= bindE /= MS_mapE /= /bindS /= fmapE /=. rewrite /cget nth_error_rcons_size /coerce. case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. From 9600b3921640150da8a38b37aa0a2a556242bca7 Mon Sep 17 00:00:00 2001 From: Takafumi Saikawa Date: Tue, 23 May 2023 18:34:07 +0900 Subject: [PATCH 42/82] shorten join'_naturality --- typed_store_model.v | 20 ++++++-------------- 1 file changed, 6 insertions(+), 14 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index ea50ad8a..03f9d0e2 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -156,7 +156,7 @@ Definition join' : [the functor of M \o M] ~~> [the functor of M] := fun _ m => bind m idfun. Set Universe Checking. -Definition actm_bind (a b c : UU0) (f : a -> b) m (g : c -> M a) : +Lemma actm_bind (a b c : UU0) (f : a -> b) m (g : c -> M a) : (actm f) (bind m g) = bind m (actm f \o g). Proof. rewrite /actm fmapE bindA. @@ -167,20 +167,12 @@ Lemma mkActoK A (m : MS _ _ A) : ofActo (mkActo m) = m. Proof. done. Qed. Unset Universe Checking. -Definition join'_naturality : naturality _ _ join'. +Lemma join'_naturality : naturality _ _ join'. Proof. -move=> a b h. -unfold join'. -rewrite /=. apply: boolp.funext => mm /=. -unfold hierarchy.actm, isFunctor.actm. -rewrite /= (actm_bind h mm idfun) /= /bind bindA. -apply f_equal. -apply f_equal. -apply boolp.funext => x /=. -rewrite /retS /= fmapE. -case: x h mm => {}a x h mm. -rewrite mkActoK /hierarchy.actm /= /actm /= fmapE !bindretf mkActoK. -reflexivity. +move=> a b h; apply: boolp.funext=> mm. +rewrite /hierarchy.actm /= actm_bind compfid. +rewrite /join' [in RHS]/bind bindA. +by congr bind. Qed. Canonical Structure naturalityTS := isNatural.Build _ _ _ join'_naturality. From d55e539e254fedd0fcfadcca9725c0f8a18cdbec Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 23 May 2023 22:30:57 +0900 Subject: [PATCH 43/82] simplify typed_store_model --- typed_store_model.v | 227 +++++++++----------------------------------- 1 file changed, 44 insertions(+), 183 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 03f9d0e2..69e329bd 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -47,176 +47,54 @@ Record binding (M : Type -> Type) := mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. Arguments mkbind {M}. +Definition M0 Env (T : UU0) := MS Env option_monad T. + #[bypass_check(positivity)] -Inductive acto : UU0 -> UU0 := - mkActo : forall T : UU0, - MS (seq (binding acto)) [the monad of option_monad] T -> acto T. +Inductive Env := mkEnv : seq (binding (M0 Env)) -> Env. + +Definition acto : UU0 -> UU0 := M0 Env. Local Notation M := acto. Local Notation coq_type := (coq_type M). -Definition ofActo T (m : M T) - : MS (seq (binding M)) [the monad of option_monad] T := - let: mkActo _ m' := m in m'. - Definition def := mkbind ml_undef (undef M). Local Notation nth_error := List.nth_error. Definition cnew T (v : coq_type T) : M (loc T) := - mkActo (fun st => let n := size st in Ret (mkloc T n, rcons st (mkbind T v))). + fun st => + let: mkEnv st := st in + let n := size st in + inr (mkloc T n, mkEnv (rcons st (mkbind T (v : coq_type T)))). -Definition coerce (T1 T2 : ml_type) (v : coq_type T1) : option (coq_type T2) := +Definition coerce T1 T2 (v : coq_type T1) : option (coq_type T2) := match ml_type_eq_dec T1 T2 with | left H => Some (eq_rect _ _ v _ H) | right _ => None end. Definition cget T (r : loc T) : M (coq_type T) := - mkActo (fun st => - if nth_error st (loc_id r) is Some (mkbind T' v) then - if coerce T v is Some u then Ret (u, st) else fail - else fail). + fun st => + let: mkEnv bs := st in + if nth_error bs (loc_id r) is Some (mkbind T' v) then + if coerce T v is Some u then inr (u, st) else inl tt + else inl tt. Definition cput T (r : loc T) (v : coq_type T) : M unit := - mkActo (fun st => - let n := loc_id r in - if nth_error st n is Some (mkbind T' _) then - if coerce T' v is Some u then - Ret (tt, set_nth def st n (mkbind T' u)) - else fail - else fail). - -Definition cchk T (r : loc T) : M unit := - mkActo (fun st => - if nth_error st (loc_id r) is Some (mkbind T' u) then - if coerce T u is Some _ then Ret (tt, st) else fail - else fail). + fun st => + let: mkEnv st := st in + let n := loc_id r in + if nth_error st n is Some (mkbind T' _) then + if coerce T' v is Some u then + let b := mkbind T' (u : coq_type _) in + inr (tt, mkEnv (set_nth def st n b)) + else inl tt + else inl tt. Definition crun (A : UU0) (m : M A) : option A := - match ofActo m nil with + match m (mkEnv nil) with | inl _ => None | inr (a, _) => Some a end. -Definition ret : forall A, idfun A -> M A := fun A a => mkActo (Ret a). -Definition bind A B (m : M A) (f : A -> M B) : M B := - mkActo (ofActo m >>= (fun a => ofActo (f a))). -Definition left_neutral : BindLaws.left_neutral bind ret. -Proof. by move=> A B a f; rewrite /bind /ret bindretf; case: (f a). Qed. -Definition right_neutral : BindLaws.right_neutral bind ret. -Proof. by move=> A a; rewrite /bind /ret bindmret; case: a. Qed. -Definition associative : BindLaws.associative bind. -Proof. by move=> A B C m f g; rewrite /bind /= bindA. Qed. - -(* Doesn't work -HB.instance Definition xyz := - isMonad_ret_bind.Build M left_neutral right_neutral associative. -*) -(* Since we couldn't build the instance, redefine notations *) -(*Local Notation "m >>= f" := (bind m f). -Local Notation "m >> f" := (bind m (fun=> f)).*) - -Definition actm (A B : UU0) (f : A -> B) (m : M A) : M B := - mkActo (actm _ _ f (ofActo m)). - -Let actm_id : FunctorLaws.id actm. -Proof. -move=> A. apply boolp.funext => -[] {A} A m. -by rewrite /actm functor_id. -Qed. - -Let actm_comp : FunctorLaws.comp actm. -Proof. -move=> A B C g h. apply boolp.funext => m. -case: m h => {A} A m h /=. -by rewrite /actm /= functor_o. -Defined. - -Lemma isFunctorTS : isFunctor.axioms_ M. -Proof. by exists actm; [exact: actm_id|exact: actm_comp]. Qed. (* NB: not Qed, o.w. ret_naturality fails *) - -Unset Universe Checking. -(*HB.instance Definition _ := isFunctor.Build acto actm_id actm_comp. fails*) -Definition tmp := isFunctor.Build acto actm_id actm_comp. -Definition functorTS : functor := HB.pack acto tmp. -Canonical Structure functorTS. -(*Canonical Structure FunctorTS : Functor M := Functor.Class isFunctorTS. -Canonical Structure functorTS : functor := Functor.Pack isFunctorTS.*) -Set Universe Checking. - -Definition ret_naturality : naturality [the functor of idfun] M ret. -Proof. by []. Qed. - -Canonical Structure naturalityTSret := isNatural.Build _ _ _ ret_naturality. - -Definition ret' : [the functor of idfun] ~> [the functor of M] := - @Nattrans.Pack _ _ ret (Nattrans.Class naturalityTSret). - -Unset Universe Checking. -Definition join' : [the functor of M \o M] ~~> [the functor of M] := - fun _ m => bind m idfun. -Set Universe Checking. - -Lemma actm_bind (a b c : UU0) (f : a -> b) m (g : c -> M a) : - (actm f) (bind m g) = bind m (actm f \o g). -Proof. -rewrite /actm fmapE bindA. -by cbv. -Qed. - -Lemma mkActoK A (m : MS _ _ A) : ofActo (mkActo m) = m. -Proof. done. Qed. - -Unset Universe Checking. -Lemma join'_naturality : naturality _ _ join'. -Proof. -move=> a b h; apply: boolp.funext=> mm. -rewrite /hierarchy.actm /= actm_bind compfid. -rewrite /join' [in RHS]/bind bindA. -by congr bind. -Qed. - -Canonical Structure naturalityTS := isNatural.Build _ _ _ join'_naturality. - -Definition join : [the functor of M \o M] ~> [the functor of M] := - @Nattrans.Pack _ _ join' (Nattrans.Class naturalityTS). - -Canonical Structure isMonadTS : isMonad.axioms_ M isFunctorTS. -exists ret' join bind. -- move=> A B f m. - rewrite /join /= /join' /bind. - apply f_equal. - rewrite bindA. - reflexivity. -- move=> A. - apply boolp.funext => x /=. - case: x => {}A m. - reflexivity. -- move=> A. - apply boolp.funext => x /=. - case: x => {}A m. - rewrite /join' /bind /=. - apply f_equal. - rewrite fmapE /ret. - rewrite bindA /=. - under [fun _ => _]boolp.funext do rewrite bindretf mkActoK. - rewrite bindmret. - reflexivity. -- move=> A /=. - rewrite /join' /bind. - apply boolp.funext => x /=. - apply f_equal. - rewrite fmapE. - rewrite !bindA. - apply f_equal. - apply boolp.funext => y /=. - rewrite bindretf mkActoK. - reflexivity. -Defined. - -Canonical Structure MonadTS : Monad M := Monad.Class isMonadTS. -Canonical Structure monadTS : monad := Monad.Pack isMonadTS. -Set Universe Checking. - (* Make ml_type an eqType *) Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. Lemma ml_type_eqP : Equality.axiom ml_type_eqb. @@ -231,12 +109,10 @@ Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. Definition cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[] st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite nth_error_rcons_size /coerce. +rewrite bindE /= bindE /= bindE /= MS_mapE /= /bindS /= fmapE /=. +rewrite /cget nth_error_rcons_size /coerce. case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. Qed. @@ -244,8 +120,7 @@ Qed. Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE bindA. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. @@ -259,8 +134,7 @@ Qed. Let cgetput T (r : loc T) (s : coq_type T) : cget r >> cput r s = cput r s. Proof. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. rewrite /coerce. @@ -277,8 +151,7 @@ Qed. Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. Proof. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE bindA /=. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. @@ -293,8 +166,7 @@ Qed. Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. @@ -311,8 +183,7 @@ Qed. Let cputget T (r: loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. case: r s k => {}T n s k /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. @@ -331,8 +202,7 @@ Qed. Let cputput T (r : loc T) (s s' : coq_type T) : cput r s >> cput r s' = cput r s'. Proof. -congr mkActo. -apply/boolp.funext => st. +apply/boolp.funext => -[st]. case: r s s' => {}T n s s' /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. @@ -350,8 +220,7 @@ Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). Proof. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. case Hr1: (nth_error _ _) => [[T1' u]|]; last first. @@ -393,10 +262,9 @@ Let cgetnewD T T' (r : loc T) (s : coq_type T') A cget r >>= (fun u => cnew s >>= fun r' => cget r >>= k r' u) = cget r >>= (fun u => cnew s >>= fun r' => k r' u u). Proof. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. +rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. rewrite {1 3}/coerce. case: ml_type_eq_dec => H /=; last by rewrite bindfailf. @@ -415,8 +283,7 @@ Let cgetnewE T1 T2 (r1 : loc T1) (s : coq_type T2) (A : UU0) cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2). Proof. move=> Hk. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] //=. @@ -430,8 +297,7 @@ Qed. Let cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip. Proof. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= 2!bindA /= /cget /cput. case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. @@ -479,8 +345,7 @@ Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. Proof. move=> H. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. case Hr1: (nth_error _ _) => [[T1' u]|]; last first. @@ -547,8 +412,7 @@ Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). Proof. move=> Hr. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite !bindA. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. @@ -587,8 +451,7 @@ Let cputnewC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A cget r >> (cnew s' >>= fun r' => cput r s >> k r') = cput r s >> (cnew s' >>= k). Proof. -congr mkActo. -apply/boolp.funext => st /=. +apply/boolp.funext => -[st] /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. @@ -611,7 +474,7 @@ Let crunret (A B : UU0) (m : M A) (s : B) : crun m -> crun (m >> Ret s) = Some s. Proof. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -by case Hm: (ofActo m [::]). +by case Hm: (m (mkEnv [::])). Qed. Let crunskip : crun skip = Some tt. @@ -621,13 +484,11 @@ Let crunnew (A : UU0) T (m : M A) (s : coq_type T) : crun m -> crun (m >> cnew s). Proof. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -by case Hm: (ofActo m [::]). +by case Hm: (m (mkEnv [::])) => [|[a [b]]]. Qed. -Unset Universe Checking. Canonical Structure isMonadTypedStoreModel := isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC cputgetC cputnewC crunret crunskip crunnew. -Set Universe Checking. End ModelTypedStore. From acb8c22083a92ad568a2b181669660d12cd11ff1 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 23 May 2023 23:02:09 +0900 Subject: [PATCH 44/82] turn option_monad into an alias --- hierarchy.v | 12 ++++++++---- monad_model.v | 23 ++++++++++++++++------- monad_transformer.v | 4 ++-- 3 files changed, 26 insertions(+), 13 deletions(-) diff --git a/hierarchy.v b/hierarchy.v index 0371db39..65b83452 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -294,11 +294,11 @@ End join_laws. End JoinLaws. HB.mixin Record isMonad (F : UU0 -> UU0) of Functor F := { - ret : FId ~> [the functor of F] ; - join : [the functor of F \o F] ~> [the functor of F] ; + ret : FId ~> F ; + join : F \o F ~> F ; bind : forall (A B : UU0), F A -> (A -> F B) -> F B ; - bindE : forall (A B : UU0) (f : A -> F B) (m : F A), - bind A B m f = join B (([the functor of F] # f) m) ; + __bindE : forall (A B : UU0) (f : A -> F B) (m : F A), + bind A B m f = join B ((F # f) m) ; joinretM : JoinLaws.left_unit ret join ; joinMret : JoinLaws.right_unit ret join ; joinA : JoinLaws.associativity join }. @@ -311,6 +311,10 @@ Notation Join := (@join _ _). Arguments bind {s A B} : simpl never. Notation "m >>= f" := (bind m f) : monae_scope. +Lemma bindE (F : monad) (A B : UU0) (f : A -> F B) (m : F A) : + m >>= f = join B ((F # f) m). +Proof. by rewrite __bindE. Qed. + Lemma eq_bind (M : monad) (A B : UU0) (m : M A) (f1 f2 : A -> M B) : f1 =1 f2 -> m >>= f1 = m >>= f2. Proof. by move=> f12; congr bind; apply boolp.funext. Qed. diff --git a/monad_model.v b/monad_model.v index e307bab5..a41c92c8 100644 --- a/monad_model.v +++ b/monad_model.v @@ -20,7 +20,7 @@ Require Import monad_transformer. (* ListMonad == list monad seq *) (* SetMonad == set monads using classical_sets *) (* ExceptMonad == exception monad E + A *) -(* option_monad == notation for ExceptMonad.acto unit *) +(* option_monad == alias for ExceptMonad.acto unit *) (* OutputMonad == output monad X * seq L *) (* EnvironmentMonad == environment monad E -> A *) (* StateMonad == state monad S -> A * S *) @@ -232,7 +232,8 @@ Lemma except_bindE (E A B : UU0) (M := ExceptMonad.acto E) m >>= f = match m with inl z => inl z | inr b => f b end. Proof. by []. Qed. -Notation option_monad := (ExceptMonad.acto unit). +Definition option_monad := ExceptMonad.acto unit. +HB.instance Definition _ := Monad.on option_monad. Module OutputMonad. Section output. @@ -799,6 +800,7 @@ Section fail. Definition option_fail : forall A, option_monad A := fun A => @throw unit A tt. Let option_bindfailf : BindLaws.left_zero (@bind _) option_fail. Proof. by []. Qed. +HB.instance Definition _ := Monad.on option_monad. HB.instance Definition _ := @isMonadFail.Build option_monad option_fail option_bindfailf. @@ -820,7 +822,7 @@ HB.export Fail. Module Except. Section except. -Let M : failMonad := [the failMonad of ExceptMonad.acto unit]. +Let M : failMonad := option_monad. Definition handle : forall A, M A -> M A -> M A := fun A m1 m2 => @handle_op unit _ (m1, (fun _ => m2)). Lemma handleE : handle = (fun A m m' => if m is inr x then m else m'). @@ -833,7 +835,7 @@ Let catchA : forall A, ssrfun.associative (@handle A). Proof. by move=> A; case. Qed. Let catchret : forall A x, @left_zero (M A) (M A) (Ret x) (@handle A). Proof. by move=> A x; case. Qed. -HB.instance Definition _ := isMonadExcept.Build (ExceptMonad.acto unit) +HB.instance Definition _ := isMonadExcept.Build option_monad catchmfail catchfailm catchA catchret. End except. End Except. @@ -1800,7 +1802,14 @@ Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. Let cnewget T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. -apply/boolp.funext => st /=. +(*apply/boolp.funext => st/=. +rewrite bindE. +transitivity ( + Join + ((stateT (seq (binding idfun)) option_monad # (fun r : loc T => cget r >>= k r)) + (cnew s)) st +) => //.*) +apply/boolp.funext => st/=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite bindE /= bindE /= bindE /= MS_mapE /= /bindS /= fmapE /=. rewrite /cget nth_error_rcons_size /coerce. @@ -2395,7 +2404,7 @@ End examples_of_algebraic_lifting. Module ModelMonadStateRun. Section modelmonadstaterun. Variable S : UU0. -Let N : monad := [the monad of ExceptMonad.acto unit]. +Let N : monad := option_monad. Definition M : stateMonad S := [the stateMonad S of stateT S N]. Let runStateT : forall A : UU0, M A -> S -> N (A * S)%type := fun A : UU0 => id. @@ -2428,7 +2437,7 @@ HB.export ModelMonadStateRun. Module ModelMonadExceptStateRun. Section modelmonadexceptstaterun. Variable S : UU0. -Let N : exceptMonad := [the exceptMonad of ExceptMonad.acto unit]. +Let N : exceptMonad := option_monad. Definition M : stateRunMonad S N := [the stateRunMonad S N of MS S N]. Definition failure : forall A, MS S N A := fun A => liftS (@fail N A). diff --git a/monad_transformer.v b/monad_transformer.v index 487b26a3..c0e32b59 100644 --- a/monad_transformer.v +++ b/monad_transformer.v @@ -145,14 +145,14 @@ by rewrite [in RHS]bindretf. Qed. HB.instance Definition _ := isMonadM_ret_bind.Build - M [the monad of MS] liftS retliftS bindliftS. + M MS liftS retliftS bindliftS. End state_monad_transformer. Definition stateT (S : UU0) := fun M => [the monad of MS S M]. HB.instance Definition _ (S : UU0) := isMonadT.Build - (stateT S) (fun M => [the monadM _ _ of @liftS S M]). + (stateT S) (@liftS S). Definition mapStateT2 (A S : UU0) (N1 N2 N3 : monad) (f : N1 (A * S)%type -> N2 (A * S)%type -> N3 (A * S)%type) From a1ef81fb1068828278276341785d0986f5f1ba91 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 24 May 2023 10:50:22 +0900 Subject: [PATCH 45/82] wip --- typed_store_model.v | 30 ++++++++++++++++++++++++------ 1 file changed, 24 insertions(+), 6 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 69e329bd..52f94619 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -51,6 +51,8 @@ Definition M0 Env (T : UU0) := MS Env option_monad T. #[bypass_check(positivity)] Inductive Env := mkEnv : seq (binding (M0 Env)) -> Env. +Definition ofEnv '(mkEnv e) := e. +Definition sizeEnv '(mkEnv e) := size e. Definition acto : UU0 -> UU0 := M0 Env. Local Notation M := acto. @@ -59,11 +61,12 @@ Local Notation coq_type := (coq_type M). Definition def := mkbind ml_undef (undef M). Local Notation nth_error := List.nth_error. +Definition extend_env T (v : coq_type T) (e : Env) := + mkEnv (rcons (ofEnv e) (mkbind T v)). +Definition fresh_loc (T : ml_type) (e : Env) := mkloc T (sizeEnv e). + Definition cnew T (v : coq_type T) : M (loc T) := - fun st => - let: mkEnv st := st in - let n := size st in - inr (mkloc T n, mkEnv (rcons st (mkbind T (v : coq_type T)))). + fun e => inr (fresh_loc T e, extend_env v e). Definition coerce T1 T2 (v : coq_type T1) : option (coq_type T2) := match ml_type_eq_dec T1 T2 with @@ -105,16 +108,31 @@ Qed. Definition ml_type_eq_mixin := EqMixin ml_type_eqP. Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. +(* WIP *) +Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M B) s : + (m >>= f) s = m s >>= uncurry f. +Proof. by []. Qed. + +Lemma cnewget0 T (s : coq_type T) A (k : loc T -> coq_type T -> M A) e : + let l := fresh_loc T e in + (cnew s >>= (fun r => cget r >>= k r)) e = (cget l >>= k l) (extend_env s e). +Proof. by case: e. Qed. + (* Prove the laws *) Definition cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. -apply/boolp.funext => -[] st /=. +apply/boolp.funext => -[st]/=. +rewrite cnewget0 !MS_bindE /=. +rewrite /cget nth_error_rcons_size /coerce. +case: ml_type_eq_dec => // H. +by rewrite -eq_rect_eq. +(*apply/boolp.funext => -[] st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite bindE /= bindE /= bindE /= MS_mapE /= /bindS /= fmapE /=. rewrite /cget nth_error_rcons_size /coerce. case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq. +by rewrite -eq_rect_eq.*) Qed. Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : From 22b9a70eaa81ec0784c2fd879de59c6078720994 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 24 May 2023 11:23:03 +0900 Subject: [PATCH 46/82] fix --- typed_store_model.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 52f94619..c21788f9 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -76,8 +76,7 @@ Definition coerce T1 T2 (v : coq_type T1) : option (coq_type T2) := Definition cget T (r : loc T) : M (coq_type T) := fun st => - let: mkEnv bs := st in - if nth_error bs (loc_id r) is Some (mkbind T' v) then + if nth_error (ofEnv st) (loc_id r) is Some (mkbind T' v) then if coerce T v is Some u then inr (u, st) else inl tt else inl tt. @@ -484,6 +483,7 @@ case: ml_type_eq_dec => // H. rewrite -eq_rect_eq. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. rewrite bindE /= /bindS /=. +rewrite /fresh_loc/=. rewrite (nth_error_size_set_nth _ _ Hr). by rewrite (nth_error_set_nth_rcons _ _ _ Hr). Qed. From df5fe684cab30b22b48f1eca6fa25ef4e68e0a0c Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 24 May 2023 12:45:45 +0900 Subject: [PATCH 47/82] wip --- typed_store_model.v | 66 ++++++++++++++++++++++++++++++++++----------- 1 file changed, 51 insertions(+), 15 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index c21788f9..86184656 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -74,6 +74,14 @@ Definition coerce T1 T2 (v : coq_type T1) : option (coq_type T2) := | right _ => None end. +Lemma coerce_sym (T T' : ml_type) (s : coq_type T) (s' : coq_type T') : + coerce T' s -> coerce T s'. +Proof. +rewrite /coerce. +case: ml_type_eq_dec => [h|h]; last by case: ml_type_eq_dec. +by case: ml_type_eq_dec => [//|g]; subst T'. +Qed. + Definition cget T (r : loc T) : M (coq_type T) := fun st => if nth_error (ofEnv st) (loc_id r) is Some (mkbind T' v) then @@ -148,22 +156,50 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. Qed. -Let cgetput T (r : loc T) (s : coq_type T) : - cget r >> cput r s = cput r s. +Let cget_Some T (r : loc T) (s : coq_type T) T' s' e A (f : _ -> _ -> M A) : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> + coerce T s' -> + (cget r >> f r s) e = f r s e. Proof. -apply/boolp.funext => -[st] /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -rewrite /coerce. -case: ml_type_eq_dec => HT; last by case: ml_type_eq_dec => // /esym //. -subst T'. -rewrite -eq_rect_eq. -case: ml_type_eq_dec => // HT. -rewrite -eq_rect_eq {HT}. -do! rewrite bindE /=. -rewrite Hr. -case: ml_type_eq_dec => // HT. -by rewrite -eq_rect_eq. +move=> H Ts'. +rewrite MS_bindE /cget H. +by case: coerce Ts'. +Qed. + +Let cget_Some' T (r : loc T) (s : coq_type T) T' s' e : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> + ~ coerce T s' -> + cget r e = fail. +Proof. by move=> H Ts'; rewrite /cget H; case: coerce Ts'. Qed. + +Let cget_None T (r : loc T) (s : coq_type T) e : + nth_error (ofEnv e) (loc_id r) = None -> + cget r e = fail. +Proof. by move=> H; rewrite /cget H. Qed. + +Let cput_Some' T (r : loc T) (s : coq_type T) T' s' e : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> + ~ coerce T s' -> + cput r s e = fail. +Proof. +move=> H Ts'; move: e H => [e] H; rewrite /cput H. +have {}Ts' : ~ coerce T' s by apply: contra_not Ts'; apply: coerce_sym. +by case: coerce Ts'. +Qed. + +Let cput_None T (r : loc T) (s : coq_type T) e : + nth_error (ofEnv e) (loc_id r) = None -> + cput r s e = fail. +Proof. by move=> H; move: e H => [e] H; rewrite /cput H. Qed. + +Let cgetput T (r : loc T) (s : coq_type T) : cget r >> cput r s = cput r s. +Proof. +apply/boolp.funext => e. +move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. + have [Ts'|Ts'] := boolp.pselect (coerce T s'). + by rewrite (cget_Some _ _ H). + by rewrite MS_bindE (cget_Some' _ H)// bindfailf (cput_Some' _ H). +by rewrite MS_bindE cget_None// bindfailf cput_None. Qed. Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. From 24c909642958f5c81edb5bc7141d2130b871f8d1 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 24 May 2023 13:55:11 +0900 Subject: [PATCH 48/82] wip --- typed_store_model.v | 64 +++++++++++++++++++++++++++++++-------------- 1 file changed, 44 insertions(+), 20 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 86184656..af24521f 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -120,9 +120,9 @@ Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M (m >>= f) s = m s >>= uncurry f. Proof. by []. Qed. -Lemma cnewget0 T (s : coq_type T) A (k : loc T -> coq_type T -> M A) e : +Lemma cnew_bind T (s : coq_type T) A (k : loc T -> coq_type T -> M A) e f : let l := fresh_loc T e in - (cnew s >>= (fun r => cget r >>= k r)) e = (cget l >>= k l) (extend_env s e). + (cnew s >>= (fun r => f r >>= k r)) e = (f l >>= k l) (extend_env s e). Proof. by case: e. Qed. (* Prove the laws *) @@ -130,7 +130,7 @@ Definition cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. apply/boolp.funext => -[st]/=. -rewrite cnewget0 !MS_bindE /=. +rewrite cnew_bind !MS_bindE /=. rewrite /cget nth_error_rcons_size /coerce. case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. @@ -160,23 +160,34 @@ Let cget_Some T (r : loc T) (s : coq_type T) T' s' e A (f : _ -> _ -> M A) : nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> coerce T s' -> (cget r >> f r s) e = f r s e. -Proof. -move=> H Ts'. -rewrite MS_bindE /cget H. -by case: coerce Ts'. -Qed. +Proof. by move=> H Ts'; rewrite MS_bindE /cget H; case: coerce Ts'. Qed. -Let cget_Some' T (r : loc T) (s : coq_type T) T' s' e : +Let cget_Some' T (r : loc T) T' s' e : nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> cget r e = fail. Proof. by move=> H Ts'; rewrite /cget H; case: coerce Ts'. Qed. -Let cget_None T (r : loc T) (s : coq_type T) e : +Let cget_None T (r : loc T) e : nth_error (ofEnv e) (loc_id r) = None -> cget r e = fail. Proof. by move=> H; rewrite /cget H. Qed. +Lemma coerce_Some (T T' : ml_type) (s : coq_type T) (h : T = T') : + coerce T' s = Some (eq_rect _ _ s _ h). +Proof. +rewrite /coerce. +case: ml_type_eq_dec => // h'. +by rewrite (boolp.Prop_irrelevance h h'). +Qed. + +Lemma coerce_Some_inv (T T' : ml_type) (s : coq_type T) : + coerce T' s -> T = T'. +Proof. +rewrite /coerce. +by case: ml_type_eq_dec. +Qed. + Let cput_Some' T (r : loc T) (s : coq_type T) T' s' e : nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> @@ -198,24 +209,37 @@ apply/boolp.funext => e. move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. have [Ts'|Ts'] := boolp.pselect (coerce T s'). by rewrite (cget_Some _ _ H). - by rewrite MS_bindE (cget_Some' _ H)// bindfailf (cput_Some' _ H). + by rewrite MS_bindE (cget_Some' H)// bindfailf (cput_Some' _ H). by rewrite MS_bindE cget_None// bindfailf cput_None. Qed. -Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. +Let cgetput_Some T (r : loc T) T' s' e : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> + coerce T s' -> + (cget r >>= cput r) e = (cget r >> skip) e. Proof. -apply/boolp.funext => -[st] /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE bindA /=. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. +move=> H Ts'; move: e H => [e] H. +rewrite !MS_bindE /cget /cput. +rewrite H. rewrite /coerce. -case: (ml_type_eq_dec T1 T) => H //. -subst T1. -rewrite bindE /= bindE /= Hr. -case: (ml_type_eq_dec T T) => H //. +case: (ml_type_eq_dec T' T) => h //. +subst T'. +rewrite bindE /= bindE/=. +rewrite H. +case: (ml_type_eq_dec T T) => // a. by rewrite -eq_rect_eq nth_error_set_nth_id. Qed. +Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. +Proof. +apply/boolp.funext => e /=. +move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. + have [Ts'|Ts'] := boolp.pselect (coerce T s'). + by rewrite (cgetput_Some H). + by rewrite !MS_bindE (cget_Some' H)// bindfailf (cget_Some' H). +by rewrite !MS_bindE cget_None// bindfailf. +Qed. + Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. From 9cf240a3b3da3960683a3bb99340b70860615f05 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 24 May 2023 15:13:42 +0900 Subject: [PATCH 49/82] wip --- typed_store_model.v | 77 +++++++++++++++++---------------------------- 1 file changed, 28 insertions(+), 49 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index af24521f..b20d2746 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -82,6 +82,12 @@ case: ml_type_eq_dec => [h|h]; last by case: ml_type_eq_dec. by case: ml_type_eq_dec => [//|g]; subst T'. Qed. +Lemma coerce_eq (T T' : ml_type) (s : coq_type T) : coerce T' s -> T = T'. +Proof. by rewrite /coerce; case: ml_type_eq_dec. Qed. + +Lemma coerce_Some (T : ml_type) (s : coq_type T) : coerce T s = Some s. +Proof. by rewrite /coerce; case: ml_type_eq_dec => // ?; rewrite -eq_rect_eq. Qed. + Definition cget T (r : loc T) : M (coq_type T) := fun st => if nth_error (ofEnv st) (loc_id r) is Some (mkbind T' v) then @@ -120,7 +126,7 @@ Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M (m >>= f) s = m s >>= uncurry f. Proof. by []. Qed. -Lemma cnew_bind T (s : coq_type T) A (k : loc T -> coq_type T -> M A) e f : +Lemma bind_cnew T (s : coq_type T) A (k : loc T -> coq_type T -> M A) e f : let l := fresh_loc T e in (cnew s >>= (fun r => f r >>= k r)) e = (f l >>= k l) (extend_env s e). Proof. by case: e. Qed. @@ -130,16 +136,10 @@ Definition cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. apply/boolp.funext => -[st]/=. -rewrite cnew_bind !MS_bindE /=. +rewrite bind_cnew !MS_bindE /=. rewrite /cget nth_error_rcons_size /coerce. case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. -(*apply/boolp.funext => -[] st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite bindE /= bindE /= bindE /= MS_mapE /= /bindS /= fmapE /=. -rewrite /cget nth_error_rcons_size /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq.*) Qed. Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : @@ -156,39 +156,24 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. Qed. -Let cget_Some T (r : loc T) (s : coq_type T) T' s' e A (f : _ -> _ -> M A) : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> - coerce T s' -> +Let Some_cget T (r : loc T) (s : coq_type T) s' e A (f : _ -> _ -> M A) : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T s') -> (cget r >> f r s) e = f r s e. -Proof. by move=> H Ts'; rewrite MS_bindE /cget H; case: coerce Ts'. Qed. +Proof. by move=> H; rewrite MS_bindE /cget H coerce_Some. Qed. +Arguments Some_cget {T r s} s'. -Let cget_Some' T (r : loc T) T' s' e : +Let nocoerce_cget T (r : loc T) T' s' e : nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> cget r e = fail. Proof. by move=> H Ts'; rewrite /cget H; case: coerce Ts'. Qed. -Let cget_None T (r : loc T) e : +Let None_cget T (r : loc T) e : nth_error (ofEnv e) (loc_id r) = None -> cget r e = fail. Proof. by move=> H; rewrite /cget H. Qed. -Lemma coerce_Some (T T' : ml_type) (s : coq_type T) (h : T = T') : - coerce T' s = Some (eq_rect _ _ s _ h). -Proof. -rewrite /coerce. -case: ml_type_eq_dec => // h'. -by rewrite (boolp.Prop_irrelevance h h'). -Qed. - -Lemma coerce_Some_inv (T T' : ml_type) (s : coq_type T) : - coerce T' s -> T = T'. -Proof. -rewrite /coerce. -by case: ml_type_eq_dec. -Qed. - -Let cput_Some' T (r : loc T) (s : coq_type T) T' s' e : +Let nocoerce_cput T (r : loc T) (s : coq_type T) T' s' e : nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> cput r s e = fail. @@ -198,7 +183,7 @@ have {}Ts' : ~ coerce T' s by apply: contra_not Ts'; apply: coerce_sym. by case: coerce Ts'. Qed. -Let cput_None T (r : loc T) (s : coq_type T) e : +Let None_cput T (r : loc T) (s : coq_type T) e : nth_error (ofEnv e) (loc_id r) = None -> cput r s e = fail. Proof. by move=> H; move: e H => [e] H; rewrite /cput H. Qed. @@ -208,26 +193,19 @@ Proof. apply/boolp.funext => e. move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. have [Ts'|Ts'] := boolp.pselect (coerce T s'). - by rewrite (cget_Some _ _ H). - by rewrite MS_bindE (cget_Some' H)// bindfailf (cput_Some' _ H). -by rewrite MS_bindE cget_None// bindfailf cput_None. + have ? := coerce_eq Ts'; subst T'. + by rewrite (@Some_cget _ _ _ s'). + by rewrite MS_bindE (nocoerce_cget H)// (nocoerce_cput _ H). +by rewrite MS_bindE None_cget// None_cput. Qed. -Let cgetput_Some T (r : loc T) T' s' e : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> - coerce T s' -> +Let cgetput_Some T (r : loc T) s' e : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T s') -> (cget r >>= cput r) e = (cget r >> skip) e. Proof. -move=> H Ts'; move: e H => [e] H. -rewrite !MS_bindE /cget /cput. -rewrite H. -rewrite /coerce. -case: (ml_type_eq_dec T' T) => h //. -subst T'. -rewrite bindE /= bindE/=. -rewrite H. -case: (ml_type_eq_dec T T) => // a. -by rewrite -eq_rect_eq nth_error_set_nth_id. +move=> H; move: e H => [e] H. +rewrite !MS_bindE /cget /cput H. +by rewrite coerce_Some {1}/bind [in LHS]/= H coerce_Some nth_error_set_nth_id. Qed. Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. @@ -235,9 +213,10 @@ Proof. apply/boolp.funext => e /=. move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. have [Ts'|Ts'] := boolp.pselect (coerce T s'). + have ? := coerce_eq Ts'; subst T'. by rewrite (cgetput_Some H). - by rewrite !MS_bindE (cget_Some' H)// bindfailf (cget_Some' H). -by rewrite !MS_bindE cget_None// bindfailf. + by rewrite !MS_bindE (nocoerce_cget H). +by rewrite !MS_bindE None_cget. Qed. Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : From 2ea483ec65f508fe21153ce8840a0c10412ee6fb Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 24 May 2023 15:39:13 +0900 Subject: [PATCH 50/82] wip --- typed_store_model.v | 33 +++++++++++++-------------------- 1 file changed, 13 insertions(+), 20 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index b20d2746..62938476 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -135,11 +135,9 @@ Proof. by case: e. Qed. Definition cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. -apply/boolp.funext => -[st]/=. +apply/boolp.funext => -[set]/=. rewrite bind_cnew !MS_bindE /=. -rewrite /cget nth_error_rcons_size /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq. +by rewrite /cget nth_error_rcons_size coerce_Some. Qed. Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : @@ -158,7 +156,7 @@ Qed. Let Some_cget T (r : loc T) (s : coq_type T) s' e A (f : _ -> _ -> M A) : nth_error (ofEnv e) (loc_id r) = Some (mkbind T s') -> - (cget r >> f r s) e = f r s e. + (cget r >>= f s) e = f s s' e. Proof. by move=> H; rewrite MS_bindE /cget H coerce_Some. Qed. Arguments Some_cget {T r s} s'. @@ -194,12 +192,12 @@ apply/boolp.funext => e. move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. have [Ts'|Ts'] := boolp.pselect (coerce T s'). have ? := coerce_eq Ts'; subst T'. - by rewrite (@Some_cget _ _ _ s'). + by rewrite (@Some_cget T r s s' e _ (fun _ _ => @cput _ r s) H). by rewrite MS_bindE (nocoerce_cget H)// (nocoerce_cput _ H). by rewrite MS_bindE None_cget// None_cput. Qed. -Let cgetput_Some T (r : loc T) s' e : +Let Some_cgetput T (r : loc T) s' e : nth_error (ofEnv e) (loc_id r) = Some (mkbind T s') -> (cget r >>= cput r) e = (cget r >> skip) e. Proof. @@ -214,7 +212,7 @@ apply/boolp.funext => e /=. move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. have [Ts'|Ts'] := boolp.pselect (coerce T s'). have ? := coerce_eq Ts'; subst T'. - by rewrite (cgetput_Some H). + by rewrite (Some_cgetput H). by rewrite !MS_bindE (nocoerce_cget H). by rewrite !MS_bindE None_cget. Qed. @@ -222,18 +220,13 @@ Qed. Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. -apply/boolp.funext => -[st] /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite /coerce. -case: (ml_type_eq_dec T1 T) => H //. -subst T1. -rewrite bindE /= bindE /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite Hr. -case: (ml_type_eq_dec T T) => H //. -by rewrite -eq_rect_eq. +apply/boolp.funext => e /=. +move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. + have [Ts'|Ts'] := boolp.pselect (coerce T s'). + have ? := coerce_eq Ts'; subst T'. + by do 3 rewrite (@Some_cget T r s' s' e)//. + by rewrite !MS_bindE (nocoerce_cget H). +by rewrite !MS_bindE None_cget. Qed. Let cputget T (r: loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : From d56690787efc904170be64ab5222813c1fcbcb94 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 24 May 2023 16:24:26 +0900 Subject: [PATCH 51/82] wip --- typed_store_model.v | 99 +++++++++++++++++++++++++-------------------- 1 file changed, 55 insertions(+), 44 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 62938476..20530f95 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -131,13 +131,30 @@ Lemma bind_cnew T (s : coq_type T) A (k : loc T -> coq_type T -> M A) e f : (cnew s >>= (fun r => f r >>= k r)) e = (f l >>= k l) (extend_env s e). Proof. by case: e. Qed. +Let Some_cget T (r : loc T) s e A (f : _ -> M A) : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T s) -> + (cget r >>= f) e = f s e. +Proof. by move=> H; rewrite MS_bindE /cget H coerce_Some. Qed. +Arguments Some_cget {T r} s. + +Let nocoerce_cget T (r : loc T) T' s' e : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> + ~ coerce T s' -> + cget r e = fail. +Proof. by move=> H Ts'; rewrite /cget H; case: coerce Ts'. Qed. + +Let None_cget T (r : loc T) e : + nth_error (ofEnv e) (loc_id r) = None -> + cget r e = fail. +Proof. by move=> H; rewrite /cget H. Qed. + (* Prove the laws *) -Definition cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : +Lemma cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. -apply/boolp.funext => -[set]/=. -rewrite bind_cnew !MS_bindE /=. -by rewrite /cget nth_error_rcons_size coerce_Some. +apply/boolp.funext => e. +rewrite bind_cnew (Some_cget s)//. +by destruct e as [e]; rewrite nth_error_rcons_size. Qed. Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : @@ -154,30 +171,13 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. Qed. -Let Some_cget T (r : loc T) (s : coq_type T) s' e A (f : _ -> _ -> M A) : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T s') -> - (cget r >>= f s) e = f s s' e. -Proof. by move=> H; rewrite MS_bindE /cget H coerce_Some. Qed. -Arguments Some_cget {T r s} s'. - -Let nocoerce_cget T (r : loc T) T' s' e : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> - ~ coerce T s' -> - cget r e = fail. -Proof. by move=> H Ts'; rewrite /cget H; case: coerce Ts'. Qed. - -Let None_cget T (r : loc T) e : - nth_error (ofEnv e) (loc_id r) = None -> - cget r e = fail. -Proof. by move=> H; rewrite /cget H. Qed. - Let nocoerce_cput T (r : loc T) (s : coq_type T) T' s' e : nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> cput r s e = fail. Proof. move=> H Ts'; move: e H => [e] H; rewrite /cput H. -have {}Ts' : ~ coerce T' s by apply: contra_not Ts'; apply: coerce_sym. +have {}Ts' : ~ coerce T' s by apply: contra_not Ts'; exact: coerce_sym. by case: coerce Ts'. Qed. @@ -192,18 +192,34 @@ apply/boolp.funext => e. move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. have [Ts'|Ts'] := boolp.pselect (coerce T s'). have ? := coerce_eq Ts'; subst T'. - by rewrite (@Some_cget T r s s' e _ (fun _ _ => @cput _ r s) H). + by rewrite (Some_cget _ _ _ _ H). by rewrite MS_bindE (nocoerce_cget H)// (nocoerce_cput _ H). by rewrite MS_bindE None_cget// None_cput. Qed. -Let Some_cgetput T (r : loc T) s' e : +Let Some_cput T (r : loc T) s e : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T s) -> + cput r s e = (@skip M) e. +Proof. +move=> H. +destruct e as [e]. +rewrite /cput/= H coerce_Some/=. +by rewrite nth_error_set_nth_id. +Qed. + +Let Some_cputget T (r : loc T) (s : coq_type T) A (k : coq_type T -> M A) + (e : Env) (s' : coq_type T) : nth_error (ofEnv e) (loc_id r) = Some (mkbind T s') -> - (cget r >>= cput r) e = (cget r >> skip) e. + (cput r s >> (cget r >>= k)) e = (cput r s >> k s) e. Proof. -move=> H; move: e H => [e] H. -rewrite !MS_bindE /cget /cput H. -by rewrite coerce_Some {1}/bind [in LHS]/= H coerce_Some nth_error_set_nth_id. +move=> H. +destruct e as [e]. +rewrite {1}/cput/= !MS_bindE H. +rewrite coerce_Some. +rewrite bindE/=. +rewrite (Some_cget s); last first. + by rewrite /= nth_error_set_nth. +by rewrite H coerce_Some. Qed. Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. @@ -212,7 +228,9 @@ apply/boolp.funext => e /=. move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. have [Ts'|Ts'] := boolp.pselect (coerce T s'). have ? := coerce_eq Ts'; subst T'. - by rewrite (Some_cgetput H). + rewrite (Some_cget _ _ _ _ H)//. + rewrite (Some_cget _ _ _ _ H)//. + by rewrite (Some_cput H). by rewrite !MS_bindE (nocoerce_cget H). by rewrite !MS_bindE None_cget. Qed. @@ -224,7 +242,7 @@ apply/boolp.funext => e /=. move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. have [Ts'|Ts'] := boolp.pselect (coerce T s'). have ? := coerce_eq Ts'; subst T'. - by do 3 rewrite (@Some_cget T r s' s' e)//. + by do 3 rewrite (@Some_cget _ _ s')//. by rewrite !MS_bindE (nocoerce_cget H). by rewrite !MS_bindE None_cget. Qed. @@ -232,20 +250,13 @@ Qed. Let cputget T (r: loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. -apply/boolp.funext => -[st] /=. -case: r s k => {}T n s k /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. -case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T'. -rewrite !bindretf /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite nth_error_set_nth. -rewrite /coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -by rewrite -eq_rect_eq !bindretf. +apply/boolp.funext => e /=. +move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. + have [Ts'|Ts'] := boolp.pselect (coerce T s'). + have ? := coerce_eq Ts'; subst T'. + by rewrite (Some_cputget _ _ H). + by rewrite !MS_bindE (nocoerce_cput _ H). +by rewrite 2!MS_bindE None_cput. Qed. Let cputput T (r : loc T) (s s' : coq_type T) : From 5a825c0aae8bb7b5fc4d144c906fdc2776d24452 Mon Sep 17 00:00:00 2001 From: Takafumi Saikawa Date: Wed, 24 May 2023 16:28:02 +0900 Subject: [PATCH 52/82] shorten cnewput --- typed_store_model.v | 28 ++++++++++++++-------------- 1 file changed, 14 insertions(+), 14 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 20530f95..ee5df45d 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -157,20 +157,6 @@ rewrite bind_cnew (Some_cget s)//. by destruct e as [e]; rewrite nth_error_rcons_size. Qed. -Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : - cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. -Proof. -apply/boolp.funext => -[st] /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE bindA. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= [in RHS]bindE /= /cput. -rewrite nth_error_rcons_size. -rewrite /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. -Qed. - Let nocoerce_cput T (r : loc T) (s : coq_type T) T' s' e : nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> @@ -186,6 +172,20 @@ Let None_cput T (r : loc T) (s : coq_type T) e : cput r s e = fail. Proof. by move=> H; move: e H => [e] H; rewrite /cput H. Qed. +Lemma bind_cnew' T (s : coq_type T) A B (k : loc T -> M A) e f : + let l := fresh_loc T e in + (cnew s >>= (fun r => bind (f r) (fun (_ : B) => k r))) e = + (f l >> k l) (extend_env s e). +Proof. by case: e. Qed. + +Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : + cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. +Proof. +apply/boolp.funext => -[st]. +rewrite bind_cnew' 2!MS_bindE /= nth_error_rcons_size coerce_Some. +by rewrite /extend_env set_nth_rcons. +Qed. + Let cgetput T (r : loc T) (s : coq_type T) : cget r >> cput r s = cput r s. Proof. apply/boolp.funext => e. From cde55dcc89ef986e439669415924c0006f5beee2 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 24 May 2023 16:36:39 +0900 Subject: [PATCH 53/82] wip --- typed_store_model.v | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/typed_store_model.v b/typed_store_model.v index ee5df45d..5549e9e9 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -260,8 +260,15 @@ by rewrite 2!MS_bindE None_cput. Qed. Let cputput T (r : loc T) (s s' : coq_type T) : - cput r s >> cput r s' = cput r s'. + cput r s >> cput r s' = cput r s'. Proof. +(*apply/boolp.funext => e. +move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T'' s'']|] H. + have [Ts''|Ts''] := boolp.pselect (coerce T s''). + have ? := coerce_eq Ts''; subst T''. + admit. + by rewrite !MS_bindE (nocoerce_cput _ H)// (nocoerce_cput _ H). +by rewrite MS_bindE !None_cput.*) apply/boolp.funext => -[st]. case: r s s' => {}T n s s' /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. From 37169221b7e5aa5f9ec65943b8f4a84573e6e144 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 24 May 2023 16:50:33 +0900 Subject: [PATCH 54/82] mv --- typed_store_model.v | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 5549e9e9..c5b56be6 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -126,11 +126,17 @@ Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M (m >>= f) s = m s >>= uncurry f. Proof. by []. Qed. -Lemma bind_cnew T (s : coq_type T) A (k : loc T -> coq_type T -> M A) e f : +Lemma bind_cnew T (s : coq_type T) A (k : loc T -> coq_type T -> M A) e (f : loc T -> M (coq_type T)): let l := fresh_loc T e in (cnew s >>= (fun r => f r >>= k r)) e = (f l >>= k l) (extend_env s e). Proof. by case: e. Qed. +Lemma bind_cnew' T (s : coq_type T) A B (k : loc T -> M A) e f : + let l := fresh_loc T e in + (cnew s >>= (fun r => f r >>= (fun (_ : B) => k r))) e = + (f l >> k l) (extend_env s e). +Proof. by case: e. Qed. + Let Some_cget T (r : loc T) s e A (f : _ -> M A) : nth_error (ofEnv e) (loc_id r) = Some (mkbind T s) -> (cget r >>= f) e = f s e. @@ -172,18 +178,12 @@ Let None_cput T (r : loc T) (s : coq_type T) e : cput r s e = fail. Proof. by move=> H; move: e H => [e] H; rewrite /cput H. Qed. -Lemma bind_cnew' T (s : coq_type T) A B (k : loc T -> M A) e f : - let l := fresh_loc T e in - (cnew s >>= (fun r => bind (f r) (fun (_ : B) => k r))) e = - (f l >> k l) (extend_env s e). -Proof. by case: e. Qed. - Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. apply/boolp.funext => -[st]. rewrite bind_cnew' 2!MS_bindE /= nth_error_rcons_size coerce_Some. -by rewrite /extend_env set_nth_rcons. +by rewrite set_nth_rcons. Qed. Let cgetput T (r : loc T) (s : coq_type T) : cget r >> cput r s = cput r s. From 6e26977f0b1a78a8cf97a682b85b82fed57b9259 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 24 May 2023 17:24:59 +0900 Subject: [PATCH 55/82] cputput --- typed_store_model.v | 31 +++++++++++++++++-------------- 1 file changed, 17 insertions(+), 14 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index c5b56be6..ca2e6d70 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -259,27 +259,30 @@ move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. by rewrite 2!MS_bindE None_cput. Qed. +Lemma Some_cputput (T : ml_type) (r : loc T) (s s' : coq_type T) + (e : Env) (s'' : coq_type T) : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T s'') -> + (cput r s >> cput r s') e = cput r s' e. +Proof. +move=> H. +destruct e as [e]. +rewrite MS_bindE/=. +rewrite H !coerce_Some. +rewrite bindE/=. +rewrite nth_error_set_nth coerce_Some/=. +by rewrite set_set_nth eqxx. +Qed. + Let cputput T (r : loc T) (s s' : coq_type T) : cput r s >> cput r s' = cput r s'. Proof. -(*apply/boolp.funext => e. +apply/boolp.funext => e. move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T'' s'']|] H. have [Ts''|Ts''] := boolp.pselect (coerce T s''). have ? := coerce_eq Ts''; subst T''. - admit. + by rewrite (Some_cputput _ _ H). by rewrite !MS_bindE (nocoerce_cput _ H)// (nocoerce_cput _ H). -by rewrite MS_bindE !None_cput.*) -apply/boolp.funext => -[st]. -case: r s s' => {}T n s s' /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. -case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T'. -rewrite !bindretf /= nth_error_set_nth /coerce. -case: ml_type_eq_dec => // H. -rewrite -eq_rect_eq. -by rewrite set_set_nth eqxx. +by rewrite MS_bindE !None_cput. Qed. Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) From eb7fad084229ec87ebcb01b08fe34264309005c4 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 25 May 2023 10:22:33 +0900 Subject: [PATCH 56/82] cgetC --- typed_store_model.v | 76 +++++++++++++++++++-------------------------- 1 file changed, 32 insertions(+), 44 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index ca2e6d70..ad65ced3 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -192,7 +192,7 @@ apply/boolp.funext => e. move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. have [Ts'|Ts'] := boolp.pselect (coerce T s'). have ? := coerce_eq Ts'; subst T'. - by rewrite (Some_cget _ _ _ _ H). + by rewrite (Some_cget s'). by rewrite MS_bindE (nocoerce_cget H)// (nocoerce_cput _ H). by rewrite MS_bindE None_cget// None_cput. Qed. @@ -225,12 +225,10 @@ Qed. Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. Proof. apply/boolp.funext => e /=. -move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. +case H : (nth_error (ofEnv e) (loc_id r)) => [[T' s']|]. have [Ts'|Ts'] := boolp.pselect (coerce T s'). have ? := coerce_eq Ts'; subst T'. - rewrite (Some_cget _ _ _ _ H)//. - rewrite (Some_cget _ _ _ _ H)//. - by rewrite (Some_cput H). + by rewrite (Some_cget s')// (Some_cget s')// (Some_cput H). by rewrite !MS_bindE (nocoerce_cget H). by rewrite !MS_bindE None_cget. Qed. @@ -239,10 +237,10 @@ Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. apply/boolp.funext => e /=. -move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. +case H : (nth_error (ofEnv e) (loc_id r)) => [[T' s']|]. have [Ts'|Ts'] := boolp.pselect (coerce T s'). have ? := coerce_eq Ts'; subst T'. - by do 3 rewrite (@Some_cget _ _ s')//. + by do 3 rewrite (Some_cget s')//. by rewrite !MS_bindE (nocoerce_cget H). by rewrite !MS_bindE None_cget. Qed. @@ -251,7 +249,7 @@ Let cputget T (r: loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. apply/boolp.funext => e /=. -move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. +case H : (nth_error (ofEnv e) (loc_id r)) => [[T' s']|]. have [Ts'|Ts'] := boolp.pselect (coerce T s'). have ? := coerce_eq Ts'; subst T'. by rewrite (Some_cputget _ _ H). @@ -277,7 +275,7 @@ Let cputput T (r : loc T) (s s' : coq_type T) : cput r s >> cput r s' = cput r s'. Proof. apply/boolp.funext => e. -move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T'' s'']|] H. +case H : (nth_error (ofEnv e) (loc_id r)) => [[T'' s'']|]. have [Ts''|Ts''] := boolp.pselect (coerce T s''). have ? := coerce_eq Ts''; subst T''. by rewrite (Some_cputput _ _ H). @@ -290,41 +288,31 @@ Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). Proof. -apply/boolp.funext => -[st] /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr1: (nth_error _ _) => [[T1' u]|]; last first. - rewrite bindfailf. - case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. - case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. - rewrite !bindE /=. - by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. -case Hr2: (nth_error _ _) => [[T2' v]|]; last first. - rewrite bindfailf. - case Hc: (coerce T1 u) => [v|]; last by rewrite bindfailf. - rewrite !bindE /=. - by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr2 bindfailf. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec _ _) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE /=. - rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr1 {1}/coerce. - case: (ml_type_eq_dec _ _) => //; by rewrite bindfailf. -subst T1'. -rewrite -eq_rect_eq. -rewrite bindE /=. -rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr2 {1}/coerce. -case: (ml_type_eq_dec T2' T2) => HT2; last by rewrite !bindfailf. -subst T2'. -rewrite -eq_rect_eq. -rewrite !bindE /=. -rewrite !bindE /=. -rewrite /bindS MS_mapE /= fmapE /= bindA /= Hr1 /coerce. -case: (ml_type_eq_dec T1 T1) => // H. -by rewrite -eq_rect_eq. +apply/boolp.funext => e /=. +case Hr1 : (nth_error (ofEnv e) (loc_id r1)) => [[T1' u]|]; last first. + rewrite MS_bindE None_cget// bindfailf. + case Hr2 : (nth_error (ofEnv e) (loc_id r2)) => [[T2' v]|]; last first. + by rewrite MS_bindE !None_cget. + have [T2v|T2v] := boolp.pselect (coerce T2 v); last first. + by rewrite MS_bindE (nocoerce_cget Hr2). + have ? := coerce_eq T2v; subst T2'. + by rewrite (Some_cget v)// MS_bindE None_cget. +have [T1u|T1u] := boolp.pselect (coerce T1 u); last first. + rewrite MS_bindE (nocoerce_cget Hr1)// bindfailf. + case Hr2 : (nth_error (ofEnv e) (loc_id r2)) => [[T2' v]|]; last first. + by rewrite MS_bindE !None_cget. + have [T2v|T2v] := boolp.pselect (coerce T2 v); last first. + by rewrite MS_bindE (nocoerce_cget Hr2). + have ? := coerce_eq T2v; subst T2'. + by rewrite (Some_cget v)// MS_bindE (nocoerce_cget Hr1). +have ? := coerce_eq T1u; subst T1'. +rewrite (Some_cget u)//. +case Hr2 : (nth_error (ofEnv e) (loc_id r2)) => [[T2' v]|]; last first. + by rewrite !MS_bindE !None_cget. +have [T2v|T2v] := boolp.pselect (coerce T2 v); last first. + by rewrite !MS_bindE (nocoerce_cget Hr2). +have ? := coerce_eq T2v; subst T2'. +by rewrite !(Some_cget v)// (Some_cget u). Qed. Let cgetnewD T T' (r : loc T) (s : coq_type T') A From f38ce64416ef0ccef8042be548d5df1c0c589f9e Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 25 May 2023 10:49:14 +0900 Subject: [PATCH 57/82] introduce spec lemmas for nth_error --- typed_store_model.v | 108 ++++++++++++++++++++++---------------------- 1 file changed, 54 insertions(+), 54 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index ad65ced3..8640a44e 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -186,15 +186,31 @@ rewrite bind_cnew' 2!MS_bindE /= nth_error_rcons_size coerce_Some. by rewrite set_nth_rcons. Qed. -Let cgetput T (r : loc T) (s : coq_type T) : cget r >> cput r s = cput r s. +Variant nth_error_spec (T : ml_type) (e : Env) (r : loc T) : Type := + | NthError : forall s', + nth_error (ofEnv e) (loc_id r) = Some (mkbind T s') -> nth_error_spec e r + | NthError_nocoerce : forall T' s', + nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> + nth_error_spec e r + | NthError_None : nth_error (ofEnv e) (loc_id r) = None -> nth_error_spec e r. + +Lemma ntherrorP (T : ml_type) (e : Env) (r : loc T) : nth_error_spec e r. Proof. -apply/boolp.funext => e. -move H : (nth_error (ofEnv e) (loc_id r)) => h; move: h H => [[T' s']|] H. +case H : (nth_error (ofEnv e) (loc_id r)) => [[T' s']|]. have [Ts'|Ts'] := boolp.pselect (coerce T s'). have ? := coerce_eq Ts'; subst T'. - by rewrite (Some_cget s'). - by rewrite MS_bindE (nocoerce_cget H)// (nocoerce_cput _ H). -by rewrite MS_bindE None_cget// None_cput. + exact: NthError H. + exact: NthError_nocoerce H Ts'. +exact: NthError_None. +Qed. + +Let cgetput T (r : loc T) (s : coq_type T) : cget r >> cput r s = cput r s. +Proof. +apply/boolp.funext => e. +have [s' H|T' s' H Ts'|H] := ntherrorP e r. +- by rewrite (Some_cget s'). +- by rewrite MS_bindE (nocoerce_cget H)// (nocoerce_cput _ H). +- by rewrite MS_bindE None_cget// None_cput. Qed. Let Some_cput T (r : loc T) s e : @@ -225,36 +241,30 @@ Qed. Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. Proof. apply/boolp.funext => e /=. -case H : (nth_error (ofEnv e) (loc_id r)) => [[T' s']|]. - have [Ts'|Ts'] := boolp.pselect (coerce T s'). - have ? := coerce_eq Ts'; subst T'. - by rewrite (Some_cget s')// (Some_cget s')// (Some_cput H). - by rewrite !MS_bindE (nocoerce_cget H). -by rewrite !MS_bindE None_cget. +have [s' H|T' s' H Ts'|H] := ntherrorP e r. +- by rewrite (Some_cget s')// (Some_cget s')// (Some_cput H). +- by rewrite !MS_bindE (nocoerce_cget H). +- by rewrite !MS_bindE None_cget. Qed. Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. apply/boolp.funext => e /=. -case H : (nth_error (ofEnv e) (loc_id r)) => [[T' s']|]. - have [Ts'|Ts'] := boolp.pselect (coerce T s'). - have ? := coerce_eq Ts'; subst T'. - by do 3 rewrite (Some_cget s')//. - by rewrite !MS_bindE (nocoerce_cget H). -by rewrite !MS_bindE None_cget. +have [s' H|T' s' H Ts'|H] := ntherrorP e r. +- by do 3 rewrite (Some_cget s')//. +- by rewrite !MS_bindE (nocoerce_cget H). +- by rewrite !MS_bindE None_cget. Qed. Let cputget T (r: loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. apply/boolp.funext => e /=. -case H : (nth_error (ofEnv e) (loc_id r)) => [[T' s']|]. - have [Ts'|Ts'] := boolp.pselect (coerce T s'). - have ? := coerce_eq Ts'; subst T'. - by rewrite (Some_cputget _ _ H). - by rewrite !MS_bindE (nocoerce_cput _ H). -by rewrite 2!MS_bindE None_cput. +have [s' H|T' s' H Ts'|H] := ntherrorP e r. +- by rewrite (Some_cputget _ _ H). +- by rewrite !MS_bindE (nocoerce_cput _ H). +- by rewrite 2!MS_bindE None_cput. Qed. Lemma Some_cputput (T : ml_type) (r : loc T) (s s' : coq_type T) @@ -275,12 +285,10 @@ Let cputput T (r : loc T) (s s' : coq_type T) : cput r s >> cput r s' = cput r s'. Proof. apply/boolp.funext => e. -case H : (nth_error (ofEnv e) (loc_id r)) => [[T'' s'']|]. - have [Ts''|Ts''] := boolp.pselect (coerce T s''). - have ? := coerce_eq Ts''; subst T''. - by rewrite (Some_cputput _ _ H). - by rewrite !MS_bindE (nocoerce_cput _ H)// (nocoerce_cput _ H). -by rewrite MS_bindE !None_cput. +have [s'' H|T'' s'' H Ts''|H] := ntherrorP e r. +- by rewrite (Some_cputput _ _ H). +- by rewrite !MS_bindE (nocoerce_cput _ H)// (nocoerce_cput _ H). +- by rewrite MS_bindE !None_cput. Qed. Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) @@ -289,30 +297,22 @@ Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). Proof. apply/boolp.funext => e /=. -case Hr1 : (nth_error (ofEnv e) (loc_id r1)) => [[T1' u]|]; last first. - rewrite MS_bindE None_cget// bindfailf. - case Hr2 : (nth_error (ofEnv e) (loc_id r2)) => [[T2' v]|]; last first. - by rewrite MS_bindE !None_cget. - have [T2v|T2v] := boolp.pselect (coerce T2 v); last first. - by rewrite MS_bindE (nocoerce_cget Hr2). - have ? := coerce_eq T2v; subst T2'. - by rewrite (Some_cget v)// MS_bindE None_cget. -have [T1u|T1u] := boolp.pselect (coerce T1 u); last first. - rewrite MS_bindE (nocoerce_cget Hr1)// bindfailf. - case Hr2 : (nth_error (ofEnv e) (loc_id r2)) => [[T2' v]|]; last first. - by rewrite MS_bindE !None_cget. - have [T2v|T2v] := boolp.pselect (coerce T2 v); last first. - by rewrite MS_bindE (nocoerce_cget Hr2). - have ? := coerce_eq T2v; subst T2'. - by rewrite (Some_cget v)// MS_bindE (nocoerce_cget Hr1). -have ? := coerce_eq T1u; subst T1'. -rewrite (Some_cget u)//. -case Hr2 : (nth_error (ofEnv e) (loc_id r2)) => [[T2' v]|]; last first. - by rewrite !MS_bindE !None_cget. -have [T2v|T2v] := boolp.pselect (coerce T2 v); last first. - by rewrite !MS_bindE (nocoerce_cget Hr2). -have ? := coerce_eq T2v; subst T2'. -by rewrite !(Some_cget v)// (Some_cget u). +have [u Hr1|u T1' Hr1 T1u|Hr1] := ntherrorP e r1. +- rewrite (Some_cget u)//. + have [v Hr2|v T2' Hr2 T2v|Hr2] := ntherrorP e r2. + + by rewrite (Some_cget v)// (Some_cget v)// (Some_cget u). + + by rewrite 2!MS_bindE (nocoerce_cget Hr2)// bindfailf. + + by rewrite !MS_bindE None_cget. +- rewrite MS_bindE (nocoerce_cget Hr1)// bindfailf. + have [v Hr2|v T2' Hr2 T2v|Hr2] := ntherrorP e r2. + + by rewrite (Some_cget v)// MS_bindE (nocoerce_cget Hr1). + + by rewrite MS_bindE (nocoerce_cget Hr2). + + by rewrite MS_bindE None_cget. +- rewrite MS_bindE None_cget// bindfailf. + have [v Hr2|v T2' Hr2 T2v|Hr2] := ntherrorP e r2. + + by rewrite (Some_cget v)// MS_bindE None_cget. + + by rewrite MS_bindE (nocoerce_cget Hr2). + + by rewrite MS_bindE !None_cget. Qed. Let cgetnewD T T' (r : loc T) (s : coq_type T') A From 522cbe9beba434f10dde391c59c9bdebd10b04a2 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 25 May 2023 11:32:52 +0900 Subject: [PATCH 58/82] cnewgetD --- typed_store_model.v | 33 +++++++++++++++++---------------- 1 file changed, 17 insertions(+), 16 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 8640a44e..a9d2bde6 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -131,10 +131,10 @@ Lemma bind_cnew T (s : coq_type T) A (k : loc T -> coq_type T -> M A) e (f : loc (cnew s >>= (fun r => f r >>= k r)) e = (f l >>= k l) (extend_env s e). Proof. by case: e. Qed. -Lemma bind_cnew' T (s : coq_type T) A B (k : loc T -> M A) e f : +Lemma bind_cnew' T (s : coq_type T) A B (k : loc T -> B -> M A) e f : let l := fresh_loc T e in - (cnew s >>= (fun r => f r >>= (fun (_ : B) => k r))) e = - (f l >> k l) (extend_env s e). + (cnew s >>= (fun r => f r >>= (fun b => k r b))) e = + (f l >>= k l) (extend_env s e). Proof. by case: e. Qed. Let Some_cget T (r : loc T) s e A (f : _ -> M A) : @@ -315,24 +315,25 @@ have [u Hr1|u T1' Hr1 T1u|Hr1] := ntherrorP e r1. + by rewrite MS_bindE !None_cget. Qed. +Let cnewgetD_helper e T T' r v (s : coq_type T') A (k : loc T' -> coq_type T -> M A) : + nth_error (ofEnv e) (loc_id r) = Some (mkbind T v) -> + (cnew s >>= (fun r' => cget r >>= k r')) e = (cnew s >>= (fun r => k r v)) e. +Proof. +move=> H. +rewrite (@bind_cnew' _ s _ _ (fun l => k l) e (fun _ => cget r))//. +by rewrite (Some_cget v) // (nth_error_rcons_some _ H). +Qed. + Let cgetnewD T T' (r : loc T) (s : coq_type T') A (k : loc T' -> coq_type T -> coq_type T -> M A) : cget r >>= (fun u => cnew s >>= fun r' => cget r >>= k r' u) = cget r >>= (fun u => cnew s >>= fun r' => k r' u u). Proof. -apply/boolp.funext => -[st] /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T1. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite (nth_error_rcons_some _ Hr). -rewrite /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq. +apply/boolp.funext => e. +have [u Hr|u T1 Hr T1u|Hr] := ntherrorP e r. +- by rewrite (Some_cget u)// (Some_cget u)// (cnewgetD_helper _ _ Hr)//. +- by rewrite !MS_bindE (nocoerce_cget Hr). +- by rewrite !MS_bindE None_cget. Qed. Let cgetnewE T1 T2 (r1 : loc T1) (s : coq_type T2) (A : UU0) From 08160fd1c0fadb51c63e2cbbe426b96bc3d066cc Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 25 May 2023 12:02:56 +0900 Subject: [PATCH 59/82] generalize bind_cnew' so that bind_cnew is not useful anymore --- typed_store_model.v | 15 +++++---------- 1 file changed, 5 insertions(+), 10 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index a9d2bde6..14936230 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -126,15 +126,9 @@ Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M (m >>= f) s = m s >>= uncurry f. Proof. by []. Qed. -Lemma bind_cnew T (s : coq_type T) A (k : loc T -> coq_type T -> M A) e (f : loc T -> M (coq_type T)): +Lemma bind_cnew T (s : coq_type T) A B (k : loc T -> B -> M A) e (f : loc T -> M B) : let l := fresh_loc T e in - (cnew s >>= (fun r => f r >>= k r)) e = (f l >>= k l) (extend_env s e). -Proof. by case: e. Qed. - -Lemma bind_cnew' T (s : coq_type T) A B (k : loc T -> B -> M A) e f : - let l := fresh_loc T e in - (cnew s >>= (fun r => f r >>= (fun b => k r b))) e = - (f l >>= k l) (extend_env s e). + (cnew s >>= (fun r => f r >>= (fun b => k r b))) e = (f l >>= k l) (extend_env s e). Proof. by case: e. Qed. Let Some_cget T (r : loc T) s e A (f : _ -> M A) : @@ -182,7 +176,7 @@ Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. apply/boolp.funext => -[st]. -rewrite bind_cnew' 2!MS_bindE /= nth_error_rcons_size coerce_Some. +rewrite bind_cnew 2!MS_bindE /= nth_error_rcons_size coerce_Some. by rewrite set_nth_rcons. Qed. @@ -315,12 +309,13 @@ have [u Hr1|u T1' Hr1 T1u|Hr1] := ntherrorP e r1. + by rewrite MS_bindE !None_cget. Qed. +(* NB: this is similar to the cnewget law *) Let cnewgetD_helper e T T' r v (s : coq_type T') A (k : loc T' -> coq_type T -> M A) : nth_error (ofEnv e) (loc_id r) = Some (mkbind T v) -> (cnew s >>= (fun r' => cget r >>= k r')) e = (cnew s >>= (fun r => k r v)) e. Proof. move=> H. -rewrite (@bind_cnew' _ s _ _ (fun l => k l) e (fun _ => cget r))//. +rewrite bind_cnew//. by rewrite (Some_cget v) // (nth_error_rcons_some _ H). Qed. From d0a6c9a7b19342696738a3477a95796016f55edd Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 25 May 2023 12:13:05 +0900 Subject: [PATCH 60/82] implicits --- typed_store_model.v | 40 ++++++++++++++++++++-------------------- 1 file changed, 20 insertions(+), 20 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 14936230..c4be9ddb 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -45,7 +45,7 @@ Import MTypedStore. Record binding (M : Type -> Type) := mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. -Arguments mkbind {M}. +Arguments mkbind {M bind_type}. Definition M0 Env (T : UU0) := MS Env option_monad T. @@ -58,11 +58,11 @@ Definition acto : UU0 -> UU0 := M0 Env. Local Notation M := acto. Local Notation coq_type := (coq_type M). -Definition def := mkbind ml_undef (undef M). +Definition def := mkbind (undef M). Local Notation nth_error := List.nth_error. Definition extend_env T (v : coq_type T) (e : Env) := - mkEnv (rcons (ofEnv e) (mkbind T v)). + mkEnv (rcons (ofEnv e) (mkbind v)). Definition fresh_loc (T : ml_type) (e : Env) := mkloc T (sizeEnv e). Definition cnew T (v : coq_type T) : M (loc T) := @@ -100,7 +100,7 @@ Definition cput T (r : loc T) (v : coq_type T) : M unit := let n := loc_id r in if nth_error st n is Some (mkbind T' _) then if coerce T' v is Some u then - let b := mkbind T' (u : coq_type _) in + let b := mkbind (u : coq_type T') in inr (tt, mkEnv (set_nth def st n b)) else inl tt else inl tt. @@ -128,17 +128,17 @@ Proof. by []. Qed. Lemma bind_cnew T (s : coq_type T) A B (k : loc T -> B -> M A) e (f : loc T -> M B) : let l := fresh_loc T e in - (cnew s >>= (fun r => f r >>= (fun b => k r b))) e = (f l >>= k l) (extend_env s e). + (cnew s >>= (fun r => f r >>= k r)) e = (f l >>= k l) (extend_env s e). Proof. by case: e. Qed. -Let Some_cget T (r : loc T) s e A (f : _ -> M A) : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T s) -> +Let Some_cget T (r : loc T) (s : coq_type T) e (A : UU0) (f : coq_type T -> M A) : + nth_error (ofEnv e) (loc_id r) = Some (mkbind s) -> (cget r >>= f) e = f s e. Proof. by move=> H; rewrite MS_bindE /cget H coerce_Some. Qed. Arguments Some_cget {T r} s. -Let nocoerce_cget T (r : loc T) T' s' e : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> +Let nocoerce_cget T (r : loc T) T' (s' : coq_type T') e : + nth_error (ofEnv e) (loc_id r) = Some (mkbind s') -> ~ coerce T s' -> cget r e = fail. Proof. by move=> H Ts'; rewrite /cget H; case: coerce Ts'. Qed. @@ -157,8 +157,8 @@ rewrite bind_cnew (Some_cget s)//. by destruct e as [e]; rewrite nth_error_rcons_size. Qed. -Let nocoerce_cput T (r : loc T) (s : coq_type T) T' s' e : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> +Let nocoerce_cput T (r : loc T) (s : coq_type T) T' (s' : coq_type T') e : + nth_error (ofEnv e) (loc_id r) = Some (mkbind s') -> ~ coerce T s' -> cput r s e = fail. Proof. @@ -181,10 +181,10 @@ by rewrite set_nth_rcons. Qed. Variant nth_error_spec (T : ml_type) (e : Env) (r : loc T) : Type := - | NthError : forall s', - nth_error (ofEnv e) (loc_id r) = Some (mkbind T s') -> nth_error_spec e r - | NthError_nocoerce : forall T' s', - nth_error (ofEnv e) (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> + | NthError : forall s : coq_type T, + nth_error (ofEnv e) (loc_id r) = Some (mkbind s) -> nth_error_spec e r + | NthError_nocoerce : forall T' (s' : coq_type T'), + nth_error (ofEnv e) (loc_id r) = Some (mkbind s') -> ~ coerce T s' -> nth_error_spec e r | NthError_None : nth_error (ofEnv e) (loc_id r) = None -> nth_error_spec e r. @@ -207,8 +207,8 @@ have [s' H|T' s' H Ts'|H] := ntherrorP e r. - by rewrite MS_bindE None_cget// None_cput. Qed. -Let Some_cput T (r : loc T) s e : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T s) -> +Let Some_cput T (r : loc T) (s : coq_type T) e : + nth_error (ofEnv e) (loc_id r) = Some (mkbind s) -> cput r s e = (@skip M) e. Proof. move=> H. @@ -219,7 +219,7 @@ Qed. Let Some_cputget T (r : loc T) (s : coq_type T) A (k : coq_type T -> M A) (e : Env) (s' : coq_type T) : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T s') -> + nth_error (ofEnv e) (loc_id r) = Some (mkbind s') -> (cput r s >> (cget r >>= k)) e = (cput r s >> k s) e. Proof. move=> H. @@ -263,7 +263,7 @@ Qed. Lemma Some_cputput (T : ml_type) (r : loc T) (s s' : coq_type T) (e : Env) (s'' : coq_type T) : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T s'') -> + nth_error (ofEnv e) (loc_id r) = Some (mkbind s'') -> (cput r s >> cput r s') e = cput r s' e. Proof. move=> H. @@ -311,7 +311,7 @@ Qed. (* NB: this is similar to the cnewget law *) Let cnewgetD_helper e T T' r v (s : coq_type T') A (k : loc T' -> coq_type T -> M A) : - nth_error (ofEnv e) (loc_id r) = Some (mkbind T v) -> + nth_error (ofEnv e) (loc_id r) = Some (mkbind v) -> (cnew s >>= (fun r' => cget r >>= k r')) e = (cnew s >>= (fun r => k r v)) e. Proof. move=> H. From 85d5b0bfc8614c73bf2ff13fd3857279f6d1d252 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 25 May 2023 12:16:57 +0900 Subject: [PATCH 61/82] more implicit --- typed_store_model.v | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index c4be9ddb..14e1330e 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -217,8 +217,8 @@ rewrite /cput/= H coerce_Some/=. by rewrite nth_error_set_nth_id. Qed. -Let Some_cputget T (r : loc T) (s : coq_type T) A (k : coq_type T -> M A) - (e : Env) (s' : coq_type T) : +Let Some_cputget T (s' s : coq_type T) (r : loc T) A (k : coq_type T -> M A) + (e : Env) : nth_error (ofEnv e) (loc_id r) = Some (mkbind s') -> (cput r s >> (cget r >>= k)) e = (cput r s >> k s) e. Proof. @@ -231,6 +231,7 @@ rewrite (Some_cget s); last first. by rewrite /= nth_error_set_nth. by rewrite H coerce_Some. Qed. +Arguments Some_cputget {T} s'. Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. Proof. @@ -256,7 +257,7 @@ Let cputget T (r: loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : Proof. apply/boolp.funext => e /=. have [s' H|T' s' H Ts'|H] := ntherrorP e r. -- by rewrite (Some_cputget _ _ H). +- by rewrite (Some_cputget s'). - by rewrite !MS_bindE (nocoerce_cput _ H). - by rewrite 2!MS_bindE None_cput. Qed. From 992569b0bdc27b285164c3a3551ae16ae96c8e4d Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Thu, 25 May 2023 13:27:50 +0900 Subject: [PATCH 62/82] simplify bind_cnew --- typed_store_model.v | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index 14e1330e..274e7563 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -126,9 +126,8 @@ Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M (m >>= f) s = m s >>= uncurry f. Proof. by []. Qed. -Lemma bind_cnew T (s : coq_type T) A B (k : loc T -> B -> M A) e (f : loc T -> M B) : - let l := fresh_loc T e in - (cnew s >>= (fun r => f r >>= k r)) e = (f l >>= k l) (extend_env s e). +Lemma bind_cnew T (s : coq_type T) A (k : loc T -> M A) e : + (cnew s >>= k) e = k (fresh_loc T e) (extend_env s e). Proof. by case: e. Qed. Let Some_cget T (r : loc T) (s : coq_type T) e (A : UU0) (f : coq_type T -> M A) : From 556d4bcdeb15c0f80a3354068c59b8e1ec293474 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 23 Jun 2023 17:29:34 +0900 Subject: [PATCH 63/82] add axioms allowing cycles in store, and example with cycle --- example_typed_store.v | 45 ++++++++++++++++++++++++++++++ hierarchy.v | 25 +++++++++++++++-- monad_model.v | 51 +++++++++++++++++++++++++++++++-- typed_store_model.v | 65 ++++++++++++++++++++++++++++++++++++++++--- 4 files changed, 177 insertions(+), 9 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 7201483e..08afb911 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -19,6 +19,7 @@ Inductive ml_type : Set := | ml_unit | ml_ref (_ : ml_type) | ml_arrow (_ : ml_type) (_ : ml_type) + | ml_rlist (_ : ml_type) | ml_undef. Inductive undef_t : Set := Undef. @@ -43,6 +44,10 @@ Local Definition loc := loc. Local Definition locT := [eqType of nat]. Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. +Inductive rlist (a : Type) (a_1 : ml_type) := + | Nil + | Cons (_ : a) (_ : loc (ml_rlist a_1)). + Section with_monad. Context [M : Type -> Type]. @@ -54,6 +59,7 @@ Fixpoint coq_type (T : ml_type) : Type := | ml_unit => unit | ml_arrow T1 T2 => coq_type T1 -> M (coq_type T2) | ml_ref T1 => loc T1 + | ml_rlist T1 => rlist (coq_type T1) T1 | ml_undef => undef_t end. End with_monad. @@ -64,6 +70,42 @@ Module IMonadTS := MonadTypedStore (MLtypes). Import MLtypes. Import IMonadTS. +Section cyclic. +Variable M : typedStoreMonad. +Notation coq_type := (@MLtypes.coq_type M). +Notation "'do' x <- m ; e" := (m >>= (fun x => e)) + (at level 60, x name, m at level 200, e at level 60). + +Definition cycle (T : ml_type) (a b : coq_type T) + : M (coq_type (ml_rlist T)) := + do r <- cnew (ml_rlist T) (Nil (coq_type T) T); + do l <- + (do v <- cnew (ml_rlist T) (Cons (coq_type T) T b r); + Ret (Cons (coq_type T) T a v)); + do _ <- cput r l; Ret l. + +Definition hd (T : ml_type) (def : coq_type T) + (param : coq_type (ml_rlist T)) : coq_type T := + match param with | Nil => def | Cons a _ => a end. + +Lemma hd_is_true : + crun (do l <- cycle ml_bool true false; Ret (hd ml_bool false l)) = Some true. +Proof. +rewrite bindA. +under eq_bind => tl. + rewrite !bindA. + under eq_bind do rewrite !bindA bindretf !bindA bindretf /=. + rewrite -bindA. + over. +rewrite -bindA crunret // -dbindA /= crungetput // bindA. +under eq_bind => tl. + rewrite !bindA. + under eq_bind do rewrite bindretf /=. + over. +by rewrite crungetnew // -(bindskipf (_ >>= _)) crunnewget // crunskip. +Qed. +End cyclic. + Section factorial. Variable M : typedStoreMonad. Notation coq_type := (@MLtypes.coq_type M). @@ -397,6 +439,7 @@ Fixpoint coq_type (T : ml_type) : Type := | ml_unit => unit | ml_arrow T1 T2 => coq_type T1 -> M (coq_type T2) | ml_ref T1 => loc T1 + | ml_rlist T1 => rlist (coq_type T1) T1 | ml_undef => undef_t end. End with_monad. @@ -465,6 +508,8 @@ Definition fact_for63 (n : coq_type ml_int) : M (coq_type ml_int) := Section fact_for63_ok. Variable n : nat. +(* Note: assuming n < max_int rather than n <= max_int is not strictly + needed, but it simplifies reasoning about loops in concrete code *) Hypothesis Hn : (Z.of_nat n < Sint63.to_Z Sint63.max_int)%Z. Let n_bounded : diff --git a/hierarchy.v b/hierarchy.v index 65b83452..8e90d05d 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -561,6 +561,17 @@ apply funext_dep => m; apply funext_dep => f. by rewrite bindE. Qed. +Lemma dbindA A B C (m : M A) (f : A -> M B) (g : A -> B -> M C) : + (m >>= fun x => f x >>= fun y => Ret (x,y)) >>= (fun xy => g xy.1 xy.2) + = (m >>= fun x => f x >>= g x). +Proof. +rewrite bindA. +apply eq_bind => x. +rewrite bindA. +apply eq_bind => y. +by rewrite bindretf. +Qed. + End monad_lemmas. Notation "'do' x <- m ; e" := (bind m (fun x => e)) (only parsing) : do_notation. @@ -1150,8 +1161,18 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) crun m -> crun (m >> Ret s) = Some s ; crunskip : crun skip = Some tt ; - crunnew : forall (A : UU0) T (m : M A) (s : coq_type M T), - crun m -> crun (m >> cnew s) ; + crunnew : forall (A : UU0) T (m : M A) (s : A -> coq_type M T), + crun m -> crun (m >>= fun x => cnew (s x)) ; + crunnewget : forall (A : UU0) T (m : M A) (s : A -> coq_type M T), + crun m -> crun (m >>= fun x => cnew (s x) >>= cget) ; + crungetnew : forall (A : UU0) T1 T2 (m : M A) (r : A -> loc T1) + (s : A -> coq_type M T2), + crun (m >>= fun x => cget (r x)) -> + crun (m >>= fun x => cnew (s x) >> cget (r x)) ; + crungetput : forall (A : UU0) T (m : M A) (r : A -> loc T) + (s : A -> coq_type M T), + crun (m >>= fun x => cget (r x)) -> + crun (m >>= fun x => cput (r x) (s x)) ; }. #[short(type=typedStoreMonad)] diff --git a/monad_model.v b/monad_model.v index a41c92c8..21a00b9b 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1767,6 +1767,14 @@ Definition coerce T1 T2 (v : coq_type M T1) : option (coq_type M T2) := | right _ => None end. +Lemma coerce_sym (T T' : ml_type) (s : coq_type M T) (s' : coq_type M T') : + coerce T' s -> coerce T s'. +Proof. +rewrite /coerce. +case: ml_type_eq_dec => [h|h]; last by case: ml_type_eq_dec. +by case: ml_type_eq_dec => [//|g]; subst T'. +Qed. + Definition cget T (r : loc T) : M (coq_type M T) := fun st => if nth_error st (loc_id r) is Some (mkbind T' v) then @@ -2180,17 +2188,54 @@ Qed. Let crunskip : crun skip = Some tt. Proof. by []. Qed. -Let crunnew (A : UU0) T (m : M A) (s : coq_type M T) : - crun m -> crun (m >> cnew s). +Let crunnew (A : UU0) T (m : M A) (s : A -> coq_type M T) : + crun m -> crun (m >>= fun x => cnew (s x)). Proof. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. by case Hm: (m [::]). Qed. +Let crunnewget (A : UU0) T (m : M A) (s : A -> coq_type M T) : + crun m -> crun (m >>= fun x => cnew (s x) >>= @cget T). +Proof. +rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +case Hm: (m [::]) => [|[a b]] // _. +by rewrite bindE /= bindE /= -(bindmret (_ >>= _)) bindA cnewget bindE. +Qed. + +Let crungetnew (A : UU0) T1 T2 (m : M A) (r : A -> loc T1) + (s : A -> coq_type M T2) : + crun (m >>= fun x => cget (r x)) -> + crun (m >>= fun x => cnew (s x) >> cget (r x)). +Proof. +rewrite /crun /= !bindE /= /bindS !MS_mapE /= !fmapE /= !bindA /=. +case Hm: (m [::]) => [|[a b]] //. +rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE /= bindA. +rewrite !bindE /= !bindE /= /cget. +case Hnth: nth_error => [[]|] //. +rewrite (nth_error_rcons_some _ Hnth). +by case Hcoe: coerce. +Qed. + +Let crungetput (A : UU0) T (m : M A) (r : A -> loc T) (s : A -> coq_type M T) : + crun (m >>= fun x => cget (r x)) -> + crun (m >>= fun x => cput (r x) (s x)). +Proof. +rewrite /crun /= !bindE /= /bindS !MS_mapE /= !fmapE /= !bindA /=. +case Hm: (m [::]) => [|[a b]] //. +rewrite !bindE /= !bindE /= /cget /cput. +case Hnth: nth_error => [[T' v]|] //. +case Hcoe: coerce => // _. +have /coerce_sym : is_true (coerce T v) by rewrite Hcoe. +move/(_ (s a)). +by case: coerce. +Qed. + Canonical Structure isMonadTypedStoreModel := isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC - cputgetC cputnewC crunret crunskip crunnew. + cputgetC cputnewC + crunret crunskip crunnew crunnewget crungetnew crungetput. (* The elpi tactic/command HB.instance failed without giving a specific error message. diff --git a/typed_store_model.v b/typed_store_model.v index 274e7563..ae5fac1d 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -179,6 +179,26 @@ rewrite bind_cnew 2!MS_bindE /= nth_error_rcons_size coerce_Some. by rewrite set_nth_rcons. Qed. +(* +Definition permutable T1 T2 (A : UU0) (k : loc T1 -> loc T2 -> M A) := + forall (n1 n2 : nat) (b1 b2 : binding M) (st : seq (binding M)), + k (mkloc T1 n1) (mkloc T2 n2) + (mkEnv (set_nth def (set_nth def st n1 b1) n2 b2)) = + k (mkloc T1 n2) (mkloc T2 n1) + (mkEnv (set_nth def (set_nth def st n2 b1) n1 b2)). + +Let cnewC T1 T2 (s : coq_type T1) (t : coq_type T2) + A (k : loc T1 -> loc T2 -> M A) : + permutable k -> + cnew s >>= (fun r => cnew t >>= k r) = + cnew t >>= (fun r => cnew s >>= k^~ r). +Proof. +move=> pmk. +apply/boolp.funext => -[st]. +rewrite !bind_cnew /fresh_loc /extend_env /= size_rcons. +Abort. +*) + Variant nth_error_spec (T : ml_type) (e : Env) (r : loc T) : Type := | NthError : forall s : coq_type T, nth_error (ofEnv e) (loc_id r) = Some (mkbind s) -> nth_error_spec e r @@ -535,15 +555,52 @@ Qed. Let crunskip : crun skip = Some tt. Proof. by []. Qed. -Let crunnew (A : UU0) T (m : M A) (s : coq_type T) : - crun m -> crun (m >> cnew s). +Let crunnew (A : UU0) T (m : M A) (s : A -> coq_type T) : + crun m -> crun (m >>= fun x => cnew (s x)). +Proof. +rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. +by case Hm: (m _). +Qed. + +Let crunnewget (A : UU0) T (m : M A) (s : A -> coq_type T) : + crun m -> crun (m >>= fun x => cnew (s x) >>= @cget T). Proof. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -by case Hm: (m (mkEnv [::])) => [|[a [b]]]. +case Hm: (m _) => [|[a b]] // _. +by rewrite bindE /= bindE /= -(bindmret (_ >>= _)) bindA cnewget bindE. +Qed. + +Let crungetnew (A : UU0) T1 T2 (m : M A) (r : A -> loc T1) + (s : A -> coq_type T2) : + crun (m >>= fun x => cget (r x)) -> + crun (m >>= fun x => cnew (s x) >> cget (r x)). +Proof. +rewrite /crun /= !bindE /= /bindS !MS_mapE /= !fmapE /= !bindA /=. +case Hm: (m _) => [|[a b]] //. +rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE /= bindA. +rewrite !bindE /= !bindE /= /cget. +case Hnth: nth_error => [[]|] //. +rewrite (nth_error_rcons_some _ Hnth). +by case Hcoe: coerce. +Qed. + +Let crungetput (A : UU0) T (m : M A) (r : A -> loc T) (s : A -> coq_type T) : + crun (m >>= fun x => cget (r x)) -> + crun (m >>= fun x => cput (r x) (s x)). +Proof. +rewrite /crun /= !bindE /= /bindS !MS_mapE /= !fmapE /= !bindA /=. +case Hm: (m _) => [|[a [b]]] //=. +rewrite !bindE /= !bindE /= /cget /cput /=. +case Hnth: nth_error => [[T' v]|] //. +case Hcoe: coerce => // _. +have /coerce_sym : is_true (coerce T v) by rewrite Hcoe. +move/(_ (s a)). +by case: coerce. Qed. Canonical Structure isMonadTypedStoreModel := isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC - cputgetC cputnewC crunret crunskip crunnew. + cputgetC cputnewC + crunret crunskip crunnew crunnewget crungetnew crungetput. End ModelTypedStore. From cc16f23346e062e5f1d77020346f37fd2029a770 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 27 Jun 2023 11:18:26 +0900 Subject: [PATCH 64/82] dbindA -> ret_uncurry? --- example_typed_store.v | 6 +++++- hierarchy.v | 8 +++----- 2 files changed, 8 insertions(+), 6 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 08afb911..158b4f78 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -97,7 +97,11 @@ under eq_bind => tl. under eq_bind do rewrite !bindA bindretf !bindA bindretf /=. rewrite -bindA. over. -rewrite -bindA crunret // -dbindA /= crungetput // bindA. +rewrite -bindA crunret //. +under eq_bind do + rewrite -(ret_uncurry (fun x => cnew (ml_rlist ml_bool) (Cons (MLtypes.coq_type ml_bool) ml_bool false x)) + (fun x x0 => cput x (Cons bool ml_bool true x0))). +rewrite -bindA /= crungetput // bindA. under eq_bind => tl. rewrite !bindA. under eq_bind do rewrite bindretf /=. diff --git a/hierarchy.v b/hierarchy.v index 8e90d05d..6ced153a 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -561,13 +561,11 @@ apply funext_dep => m; apply funext_dep => f. by rewrite bindE. Qed. -Lemma dbindA A B C (m : M A) (f : A -> M B) (g : A -> B -> M C) : - (m >>= fun x => f x >>= fun y => Ret (x,y)) >>= (fun xy => g xy.1 xy.2) - = (m >>= fun x => f x >>= g x). +Lemma ret_uncurry A B C (f : A -> M B) (g : A -> B -> M C) x : + (f x >>= fun y => Ret (x,y)) >>= (fun xy => g xy.1 xy.2) + = (f x >>= g x). Proof. rewrite bindA. -apply eq_bind => x. -rewrite bindA. apply eq_bind => y. by rewrite bindretf. Qed. From e94438d2d773d7288772c9c201651dd6954dcc28 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 27 Jun 2023 12:08:29 +0900 Subject: [PATCH 65/82] check instance of typed_store_model --- example_typed_store.v | 15 ++++++++------- hierarchy.v | 2 +- monad_model.v | 3 +-- typed_store_model.v | 5 +---- 4 files changed, 11 insertions(+), 14 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 158b4f78..69910e38 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -24,8 +24,6 @@ Inductive ml_type : Set := Inductive undef_t : Set := Undef. -Variant loc : ml_type -> Type := mkloc T : nat -> loc T. - Module MLtypes. Definition ml_type_eq_dec (T1 T2 : ml_type) : {T1=T2}+{T1<>T2}. revert T2; induction T1; destruct T2; @@ -39,8 +37,8 @@ revert T2; induction T1; destruct T2; Defined. Local Definition ml_type := ml_type. -Local Definition undef := Undef. -Local Definition loc := loc. +Local Definition undef (M : UU0 -> UU0) := Undef. +Variant loc : ml_type -> Type := mkloc T : nat -> loc T. Local Definition locT := [eqType of nat]. Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. @@ -70,6 +68,9 @@ Module IMonadTS := MonadTypedStore (MLtypes). Import MLtypes. Import IMonadTS. +Require Import typed_store_model. +Module ModelTS := ModelTypedStore (MLtypes). + Section cyclic. Variable M : typedStoreMonad. Notation coq_type := (@MLtypes.coq_type M). @@ -427,10 +428,10 @@ End Int63. Module MLtypes63. Local Definition ml_type_eq_dec := ml_type_eq_dec. Local Definition ml_type := ml_type. -Local Definition undef := Undef. -Local Definition loc := loc. +Local Definition undef (M : UU0 -> UU0) := Undef. +Variant loc : ml_type -> Set := mkloc T : nat -> loc T. Local Definition locT := [eqType of nat]. -Local Definition loc_id := @loc_id. +Local Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. Section with_monad. Context [M : Type -> Type]. diff --git a/hierarchy.v b/hierarchy.v index 6ced153a..318d95a8 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1093,7 +1093,7 @@ HB.structure Definition MonadPlusArray (S : UU0) (I : eqType) := Module Type MLTY. Parameter ml_type : Set. Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Parameter loc : ml_type -> Type. +Variant loc : ml_type -> Type := mkloc T : nat -> loc T. Parameter locT : eqType. Parameter loc_id : forall T, loc T -> locT. Parameter coq_type : forall M : Type -> Type, ml_type -> Type. diff --git a/monad_model.v b/monad_model.v index 21a00b9b..8bbc03f6 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1720,7 +1720,6 @@ Require Import JMeq. Module Type MLTYweak. Parameter ml_type : Set. Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Variant loc : ml_type -> Set := mkloc T : nat -> loc T. Parameter coq_type0 : ml_type -> Type. Parameter ml_undef : ml_type. Parameter undef : coq_type0 ml_undef. @@ -1733,7 +1732,7 @@ Module MLtypes'. Definition ml_type := ml_type. Definition ml_type_eq_dec := ml_type_eq_dec. Definition coq_type (M : UU0 -> UU0) := coq_type0. -Definition loc := loc. +Variant loc : ml_type -> Type := mkloc T : nat -> loc T. Definition locT := [eqType of nat]. Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. End MLtypes'. diff --git a/typed_store_model.v b/typed_store_model.v index ae5fac1d..64cf84a2 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -31,10 +31,7 @@ End MLTYdef. Module ModelTypedStore (MLtypes : MLTYdef). Import MLtypes. Module MLtypes'. -Definition ml_type := ml_type. -Definition ml_type_eq_dec := ml_type_eq_dec. -Definition coq_type := coq_type. -Definition loc := loc. +Include MLtypes. Definition locT := [eqType of nat]. Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. End MLtypes'. From 5be44137a808c0ea7b58282a186b691844d896d3 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 27 Jun 2023 12:19:37 +0900 Subject: [PATCH 66/82] typo --- example_typed_store.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/example_typed_store.v b/example_typed_store.v index 69910e38..7e0973b9 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -100,7 +100,7 @@ under eq_bind => tl. over. rewrite -bindA crunret //. under eq_bind do - rewrite -(ret_uncurry (fun x => cnew (ml_rlist ml_bool) (Cons (MLtypes.coq_type ml_bool) ml_bool false x)) + rewrite -(ret_uncurry (fun x => cnew (ml_rlist ml_bool) (Cons (MLtypes.coq_type _ ml_bool) ml_bool false x)) (fun x x0 => cput x (Cons bool ml_bool true x0))). rewrite -bindA /= crungetput // bindA. under eq_bind => tl. From 263fecf346f60d8c6423b08f65d756628ae5ba69 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 27 Jun 2023 12:26:21 +0900 Subject: [PATCH 67/82] bindA_uncurry --- example_typed_store.v | 8 ++------ hierarchy.v | 9 --------- monad_lib.v | 10 ++++++++++ 3 files changed, 12 insertions(+), 15 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 7e0973b9..89f7cb04 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -4,7 +4,7 @@ Require Import ZArith. From mathcomp Require Import all_ssreflect. From mathcomp Require boolp. From infotheo Require Import ssrZ. -Require Import monae_lib hierarchy. (*monad_lib fail_lib state_lib.*) +Require Import monae_lib hierarchy monad_lib. (*fail_lib state_lib.*) (******************************************************************************) (* Type store examples *) @@ -98,11 +98,7 @@ under eq_bind => tl. under eq_bind do rewrite !bindA bindretf !bindA bindretf /=. rewrite -bindA. over. -rewrite -bindA crunret //. -under eq_bind do - rewrite -(ret_uncurry (fun x => cnew (ml_rlist ml_bool) (Cons (MLtypes.coq_type _ ml_bool) ml_bool false x)) - (fun x x0 => cput x (Cons bool ml_bool true x0))). -rewrite -bindA /= crungetput // bindA. +rewrite -bindA crunret // -bindA_uncurry /= crungetput // bindA. under eq_bind => tl. rewrite !bindA. under eq_bind do rewrite bindretf /=. diff --git a/hierarchy.v b/hierarchy.v index 318d95a8..ada8e05f 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -561,15 +561,6 @@ apply funext_dep => m; apply funext_dep => f. by rewrite bindE. Qed. -Lemma ret_uncurry A B C (f : A -> M B) (g : A -> B -> M C) x : - (f x >>= fun y => Ret (x,y)) >>= (fun xy => g xy.1 xy.2) - = (f x >>= g x). -Proof. -rewrite bindA. -apply eq_bind => y. -by rewrite bindretf. -Qed. - End monad_lemmas. Notation "'do' x <- m ; e" := (bind m (fun x => e)) (only parsing) : do_notation. diff --git a/monad_lib.v b/monad_lib.v index 1b504a5d..fb22097b 100644 --- a/monad_lib.v +++ b/monad_lib.v @@ -153,6 +153,16 @@ HB.instance Definition _ X := Definition uncurry_F X : functor := [the functor of @uncurry_M X]. End uncurry_functor. +Lemma bind_uncurry (M : monad) A B C (f : A -> M B) (g : A -> B -> M C) x : + (f x >>= fun y => Ret (x,y)) >>= (fun xy => g xy.1 xy.2)(*uncurry g*) = + (f x >>= g x). +Proof. by rewrite bindA; under eq_bind do rewrite bindretf. Qed. + +Lemma bindA_uncurry (M : monad) A B C (m : M A) (f : A -> M B) (g : A -> B -> M C) : + (m >>= fun x => f x >>= fun y => Ret (x,y)) >>= (fun xy => g xy.1 xy.2) = + (m >>= fun x => f x >>= g x). +Proof. by rewrite bindA; by under eq_bind do rewrite bind_uncurry. Qed. + Section exponential_functor. Variable A : UU0. Definition exponential_M (X : UU0) := A -> X. From f3e9f145f2dfc3ecba0797d1cb9816e0f5037e34 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 27 Jun 2023 13:00:28 +0900 Subject: [PATCH 68/82] build actual model in example_typed_store --- example_typed_store.v | 16 ++++++++++++++++ typed_store_model.v | 2 +- 2 files changed, 17 insertions(+), 1 deletion(-) diff --git a/example_typed_store.v b/example_typed_store.v index 89f7cb04..fe696d79 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -68,9 +68,25 @@ Module IMonadTS := MonadTypedStore (MLtypes). Import MLtypes. Import IMonadTS. +Module Model. Require Import typed_store_model. Module ModelTS := ModelTypedStore (MLtypes). +Require Import fail_lib state_lib trace_lib. +Require Import monad_transformer monad_model typed_store_model. +From HB Require Import structures. +Import ModelTS. +Check [the monad of MS Env option_monad]. +(* Fails +Check [the monad of acto]. +Definition D := + isMonadTypedStore.Build _ cnewget cnewput cgetput cgetputskip + cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC + cputgetC cputnewC + crunret crunskip crunnew crunnewget crungetnew crungetput. + *) +End Model. + Section cyclic. Variable M : typedStoreMonad. Notation coq_type := (@MLtypes.coq_type M). diff --git a/typed_store_model.v b/typed_store_model.v index 64cf84a2..20e4b9c9 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -595,7 +595,7 @@ move/(_ (s a)). by case: coerce. Qed. -Canonical Structure isMonadTypedStoreModel := +Definition isMonadTypedStoreModel := isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC cputgetC cputnewC From a4a1cfe61fdcfac71ce892b8b51dc5b820e8e639 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 27 Jun 2023 13:15:43 +0900 Subject: [PATCH 69/82] can define outside typed_store_model, but HB.instance still fails --- example_typed_store.v | 19 +++++++------- typed_store_model.v | 58 +++++++++++++++++++++---------------------- 2 files changed, 39 insertions(+), 38 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index fe696d79..cb10226c 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -64,10 +64,6 @@ End with_monad. Local Definition ml_undef := ml_undef. End MLtypes. -Module IMonadTS := MonadTypedStore (MLtypes). -Import MLtypes. -Import IMonadTS. - Module Model. Require Import typed_store_model. Module ModelTS := ModelTypedStore (MLtypes). @@ -76,17 +72,22 @@ Require Import fail_lib state_lib trace_lib. Require Import monad_transformer monad_model typed_store_model. From HB Require Import structures. Import ModelTS. -Check [the monad of MS Env option_monad]. -(* Fails -Check [the monad of acto]. Definition D := - isMonadTypedStore.Build _ cnewget cnewput cgetput cgetputskip + MTypedStore.isMonadTypedStore.Build _ cnewget cnewput cgetput cgetputskip + cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC + cputgetC cputnewC + crunret crunskip crunnew crunnewget crungetnew crungetput. +Fail HB.instance Definition _ := + MTypedStore.isMonadTypedStore.Build _ cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC cputgetC cputnewC crunret crunskip crunnew crunnewget crungetnew crungetput. - *) End Model. +Module IMonadTS := MonadTypedStore (MLtypes). +Import MLtypes. +Import IMonadTS. + Section cyclic. Variable M : typedStoreMonad. Notation coq_type := (@MLtypes.coq_type M). diff --git a/typed_store_model.v b/typed_store_model.v index 20e4b9c9..a0d7772d 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -51,7 +51,7 @@ Inductive Env := mkEnv : seq (binding (M0 Env)) -> Env. Definition ofEnv '(mkEnv e) := e. Definition sizeEnv '(mkEnv e) := size e. -Definition acto : UU0 -> UU0 := M0 Env. +Definition acto : UU0 -> UU0 := MS Env option_monad. Local Notation M := acto. Local Notation coq_type := (coq_type M). @@ -127,19 +127,19 @@ Lemma bind_cnew T (s : coq_type T) A (k : loc T -> M A) e : (cnew s >>= k) e = k (fresh_loc T e) (extend_env s e). Proof. by case: e. Qed. -Let Some_cget T (r : loc T) (s : coq_type T) e (A : UU0) (f : coq_type T -> M A) : +Lemma Some_cget T (r : loc T) (s : coq_type T) e (A : UU0) (f : coq_type T -> M A) : nth_error (ofEnv e) (loc_id r) = Some (mkbind s) -> (cget r >>= f) e = f s e. Proof. by move=> H; rewrite MS_bindE /cget H coerce_Some. Qed. Arguments Some_cget {T r} s. -Let nocoerce_cget T (r : loc T) T' (s' : coq_type T') e : +Lemma nocoerce_cget T (r : loc T) T' (s' : coq_type T') e : nth_error (ofEnv e) (loc_id r) = Some (mkbind s') -> ~ coerce T s' -> cget r e = fail. Proof. by move=> H Ts'; rewrite /cget H; case: coerce Ts'. Qed. -Let None_cget T (r : loc T) e : +Lemma None_cget T (r : loc T) e : nth_error (ofEnv e) (loc_id r) = None -> cget r e = fail. Proof. by move=> H; rewrite /cget H. Qed. @@ -153,7 +153,7 @@ rewrite bind_cnew (Some_cget s)//. by destruct e as [e]; rewrite nth_error_rcons_size. Qed. -Let nocoerce_cput T (r : loc T) (s : coq_type T) T' (s' : coq_type T') e : +Lemma nocoerce_cput T (r : loc T) (s : coq_type T) T' (s' : coq_type T') e : nth_error (ofEnv e) (loc_id r) = Some (mkbind s') -> ~ coerce T s' -> cput r s e = fail. @@ -163,12 +163,12 @@ have {}Ts' : ~ coerce T' s by apply: contra_not Ts'; exact: coerce_sym. by case: coerce Ts'. Qed. -Let None_cput T (r : loc T) (s : coq_type T) e : +Lemma None_cput T (r : loc T) (s : coq_type T) e : nth_error (ofEnv e) (loc_id r) = None -> cput r s e = fail. Proof. by move=> H; move: e H => [e] H; rewrite /cput H. Qed. -Let cnewput T (s t : coq_type T) A (k : loc T -> M A) : +Lemma cnewput T (s t : coq_type T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. apply/boolp.funext => -[st]. @@ -184,7 +184,7 @@ Definition permutable T1 T2 (A : UU0) (k : loc T1 -> loc T2 -> M A) := k (mkloc T1 n2) (mkloc T2 n1) (mkEnv (set_nth def (set_nth def st n2 b1) n1 b2)). -Let cnewC T1 T2 (s : coq_type T1) (t : coq_type T2) +Lemma cnewC T1 T2 (s : coq_type T1) (t : coq_type T2) A (k : loc T1 -> loc T2 -> M A) : permutable k -> cnew s >>= (fun r => cnew t >>= k r) = @@ -214,7 +214,7 @@ case H : (nth_error (ofEnv e) (loc_id r)) => [[T' s']|]. exact: NthError_None. Qed. -Let cgetput T (r : loc T) (s : coq_type T) : cget r >> cput r s = cput r s. +Lemma cgetput T (r : loc T) (s : coq_type T) : cget r >> cput r s = cput r s. Proof. apply/boolp.funext => e. have [s' H|T' s' H Ts'|H] := ntherrorP e r. @@ -223,7 +223,7 @@ have [s' H|T' s' H Ts'|H] := ntherrorP e r. - by rewrite MS_bindE None_cget// None_cput. Qed. -Let Some_cput T (r : loc T) (s : coq_type T) e : +Lemma Some_cput T (r : loc T) (s : coq_type T) e : nth_error (ofEnv e) (loc_id r) = Some (mkbind s) -> cput r s e = (@skip M) e. Proof. @@ -233,7 +233,7 @@ rewrite /cput/= H coerce_Some/=. by rewrite nth_error_set_nth_id. Qed. -Let Some_cputget T (s' s : coq_type T) (r : loc T) A (k : coq_type T -> M A) +Lemma Some_cputget T (s' s : coq_type T) (r : loc T) A (k : coq_type T -> M A) (e : Env) : nth_error (ofEnv e) (loc_id r) = Some (mkbind s') -> (cput r s >> (cget r >>= k)) e = (cput r s >> k s) e. @@ -249,7 +249,7 @@ by rewrite H coerce_Some. Qed. Arguments Some_cputget {T} s'. -Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. +Lemma cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. Proof. apply/boolp.funext => e /=. have [s' H|T' s' H Ts'|H] := ntherrorP e r. @@ -258,7 +258,7 @@ have [s' H|T' s' H Ts'|H] := ntherrorP e r. - by rewrite !MS_bindE None_cget. Qed. -Let cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : +Lemma cgetget T (r : loc T) (A : UU0) (k : coq_type T -> coq_type T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. apply/boolp.funext => e /=. @@ -268,7 +268,7 @@ have [s' H|T' s' H Ts'|H] := ntherrorP e r. - by rewrite !MS_bindE None_cget. Qed. -Let cputget T (r: loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : +Lemma cputget T (r: loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. apply/boolp.funext => e /=. @@ -292,7 +292,7 @@ rewrite nth_error_set_nth coerce_Some/=. by rewrite set_set_nth eqxx. Qed. -Let cputput T (r : loc T) (s s' : coq_type T) : +Lemma cputput T (r : loc T) (s s' : coq_type T) : cput r s >> cput r s' = cput r s'. Proof. apply/boolp.funext => e. @@ -302,7 +302,7 @@ have [s'' H|T'' s'' H Ts''|H] := ntherrorP e r. - by rewrite MS_bindE !None_cput. Qed. -Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) +Lemma cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) (k : coq_type T1 -> coq_type T2 -> M A) : cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). @@ -327,7 +327,7 @@ have [u Hr1|u T1' Hr1 T1u|Hr1] := ntherrorP e r1. Qed. (* NB: this is similar to the cnewget law *) -Let cnewgetD_helper e T T' r v (s : coq_type T') A (k : loc T' -> coq_type T -> M A) : +Lemma cnewgetD_helper e T T' r v (s : coq_type T') A (k : loc T' -> coq_type T -> M A) : nth_error (ofEnv e) (loc_id r) = Some (mkbind v) -> (cnew s >>= (fun r' => cget r >>= k r')) e = (cnew s >>= (fun r => k r v)) e. Proof. @@ -336,7 +336,7 @@ rewrite bind_cnew//. by rewrite (Some_cget v) // (nth_error_rcons_some _ H). Qed. -Let cgetnewD T T' (r : loc T) (s : coq_type T') A +Lemma cgetnewD T T' (r : loc T) (s : coq_type T') A (k : loc T' -> coq_type T -> coq_type T -> M A) : cget r >>= (fun u => cnew s >>= fun r' => cget r >>= k r' u) = cget r >>= (fun u => cnew s >>= fun r' => k r' u u). @@ -348,7 +348,7 @@ have [u Hr|u T1 Hr T1u|Hr] := ntherrorP e r. - by rewrite !MS_bindE None_cget. Qed. -Let cgetnewE T1 T2 (r1 : loc T1) (s : coq_type T2) (A : UU0) +Lemma cgetnewE T1 T2 (r1 : loc T1) (s : coq_type T2) (A : UU0) (k1 k2 : loc T2 -> M A) : (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2). @@ -365,7 +365,7 @@ rewrite fmapE /= bindA /= 2!bindE /=. by rewrite Hk // neq_ltn /= (nth_error_size Hr1). Qed. -Let cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : +Lemma cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip. Proof. apply/boolp.funext => -[st] /=. @@ -410,7 +410,7 @@ rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. by case: ml_type_eq_dec. Qed. -Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) +Lemma cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) (s2 : coq_type T2) (A : UU0) : loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. @@ -477,7 +477,7 @@ case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq set_set_nth (negbTE Hr). Qed. -Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) +Lemma cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) (A : UU0) (k : coq_type T2 -> M A) : loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). @@ -517,7 +517,7 @@ case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq bindE. Qed. -Let cputnewC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A +Lemma cputnewC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A (k : loc T' -> M A) : cget r >> (cnew s' >>= fun r' => cput r s >> k r') = cput r s >> (cnew s' >>= k). @@ -542,24 +542,24 @@ rewrite (nth_error_size_set_nth _ _ Hr). by rewrite (nth_error_set_nth_rcons _ _ _ Hr). Qed. -Let crunret (A B : UU0) (m : M A) (s : B) : +Lemma crunret (A B : UU0) (m : M A) (s : B) : crun m -> crun (m >> Ret s) = Some s. Proof. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. by case Hm: (m (mkEnv [::])). Qed. -Let crunskip : crun skip = Some tt. +Lemma crunskip : crun skip = Some tt. Proof. by []. Qed. -Let crunnew (A : UU0) T (m : M A) (s : A -> coq_type T) : +Lemma crunnew (A : UU0) T (m : M A) (s : A -> coq_type T) : crun m -> crun (m >>= fun x => cnew (s x)). Proof. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. by case Hm: (m _). Qed. -Let crunnewget (A : UU0) T (m : M A) (s : A -> coq_type T) : +Lemma crunnewget (A : UU0) T (m : M A) (s : A -> coq_type T) : crun m -> crun (m >>= fun x => cnew (s x) >>= @cget T). Proof. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. @@ -567,7 +567,7 @@ case Hm: (m _) => [|[a b]] // _. by rewrite bindE /= bindE /= -(bindmret (_ >>= _)) bindA cnewget bindE. Qed. -Let crungetnew (A : UU0) T1 T2 (m : M A) (r : A -> loc T1) +Lemma crungetnew (A : UU0) T1 T2 (m : M A) (r : A -> loc T1) (s : A -> coq_type T2) : crun (m >>= fun x => cget (r x)) -> crun (m >>= fun x => cnew (s x) >> cget (r x)). @@ -581,7 +581,7 @@ rewrite (nth_error_rcons_some _ Hnth). by case Hcoe: coerce. Qed. -Let crungetput (A : UU0) T (m : M A) (r : A -> loc T) (s : A -> coq_type T) : +Lemma crungetput (A : UU0) T (m : M A) (r : A -> loc T) (s : A -> coq_type T) : crun (m >>= fun x => cget (r x)) -> crun (m >>= fun x => cput (r x) (s x)). Proof. From d92b5df00e393d40e97faa14f92b42a9ddad6be3 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 27 Jun 2023 14:38:03 +0900 Subject: [PATCH 70/82] cgetnewE --- monad_model.v | 16 +++++++--------- typed_store_model.v | 15 ++++++--------- 2 files changed, 13 insertions(+), 18 deletions(-) diff --git a/monad_model.v b/monad_model.v index 8bbc03f6..16373ead 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1664,16 +1664,14 @@ Context [T : Type] (def : T) (st : seq T). (* Basic lemmas for standard library's nth_error *) Local Notation nth_error := List.nth_error. -Lemma nth_error_set_nth n x : - nth_error (set_nth def st n x) n = Some x. +Lemma nth_error_set_nth n x : nth_error (set_nth def st n x) n = Some x. Proof. elim: n st => [|z IH] [] // {IH}. elim: z.+1 => [|n <-] //=. by rewrite set_nth_nil. Qed. -Lemma nth_error_rcons_size b : - nth_error (rcons st b) (size st) = Some b. +Lemma nth_error_rcons_size b : nth_error (rcons st b) (size st) = Some b. Proof. by elim: st. Qed. Lemma nth_error_rcons_some n a b : @@ -1688,7 +1686,8 @@ Lemma nth_error_set_nth_other m n a b : m != n -> nth_error st m = Some a -> nth_error (set_nth def st n b) m = Some a. -Proof. elim: m st n => [|m IH] [|c st'] [|n] //=; rewrite eqSS => *; exact: IH. +Proof. +elim: m st n => [|m IH] [|c st'] [|n] //=; rewrite eqSS => *; exact: IH. Qed. Lemma nth_error_set_nth_none m n a b : @@ -1697,16 +1696,14 @@ Lemma nth_error_set_nth_none m n a b : nth_error (set_nth def st n b) m = None. Proof. by elim: m st n => [|m IH] [|c st'] [|n] //=; apply IH. Qed. -Lemma nth_error_size n a : - nth_error st n = Some a -> n < size st. +Lemma nth_error_size n a : nth_error st n = Some a -> n < size st. Proof. by elim: n st => [|n IH] [|c st'] //= /IH. Qed. Lemma nth_error_size_set_nth n a b : nth_error st n = Some a -> size (set_nth def st n b) = size st. Proof. by rewrite size_set_nth => /nth_error_size /maxn_idPr. Qed. -Lemma set_nth_rcons a b : - set_nth def (rcons st a) (size st) b = rcons st b. +Lemma set_nth_rcons a b : set_nth def (rcons st a) (size st) b = rcons st b. Proof. by elim: st => //= c st' ->. Qed. Lemma nth_error_set_nth_rcons n a b c : @@ -1714,6 +1711,7 @@ Lemma nth_error_set_nth_rcons n a b c : set_nth def (rcons st c) n b = rcons (set_nth def st n b) c. Proof. by elim: n st => [|n IH] [|d st'] //= /IH ->. Qed. End nth_error. +Arguments nth_error_size {T st n a}. Require Import JMeq. diff --git a/typed_store_model.v b/typed_store_model.v index a0d7772d..94fdef3c 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -354,15 +354,12 @@ Lemma cgetnewE T1 T2 (r1 : loc T1) (s : coq_type T2) (A : UU0) cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2). Proof. move=> Hk. -apply/boolp.funext => -[st] /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] //=. -case Hc: (coerce T1 v1) => [u1|] //. -rewrite bindE /= bindE /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= [in RHS]bindE /= [in RHS]bindE /= /bindS MS_mapE /=. -rewrite fmapE /= bindA /= 2!bindE /=. -by rewrite Hk // neq_ltn /= (nth_error_size Hr1). +apply/boolp.funext => e. +have [u r1u|T' s' Hr1 T1s'|Hr] := ntherrorP e r1. +- rewrite (Some_cget u)// (Some_cget u)// !bind_cnew Hk// neq_ltn. + by case: e r1u => /= e /nth_error_size ->. +- by rewrite !MS_bindE (nocoerce_cget Hr1). +- by rewrite !MS_bindE None_cget. Qed. Lemma cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : From 536b1083d1e636f9e023d8d08362af537c41629b Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 30 Jun 2023 18:44:00 +0900 Subject: [PATCH 71/82] check examples with both models --- example_typed_store.v | 43 +++++++++++++++++++++++++---------- monad_model.v | 52 ++++++++++++++++++++++--------------------- 2 files changed, 58 insertions(+), 37 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index cb10226c..93bf2319 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -61,42 +61,61 @@ Fixpoint coq_type (T : ml_type) : Type := | ml_undef => undef_t end. End with_monad. +Local Definition coq_type0 := @coq_type idfun. Local Definition ml_undef := ml_undef. End MLtypes. -Module Model. -Require Import typed_store_model. -Module ModelTS := ModelTypedStore (MLtypes). +(* Check that the models can indeed be instanciated, for soundness *) +Module Model0. +Require monad_model. +(* This model does not allow functions in the store, but can be defined + without endangering soundness. *) +Module ModelTS := monad_model.ModelTypedStore (MLtypes). +(* Attempt to build an instance with HB *) Require Import fail_lib state_lib trace_lib. -Require Import monad_transformer monad_model typed_store_model. -From HB Require Import structures. +Require Import monad_transformer monad_model. Import ModelTS. -Definition D := +From HB Require Import structures. +(* This fails *) +Fail HB.instance Definition _ := MTypedStore.isMonadTypedStore.Build _ cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC cputgetC cputnewC crunret crunskip crunnew crunnewget crungetnew crungetput. -Fail HB.instance Definition _ := +(* But we can still check that the definition is correct (again) *) +Definition isMonadTypedStoreModel := MTypedStore.isMonadTypedStore.Build _ cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC cputgetC cputnewC crunret crunskip crunnew crunnewget crungetnew crungetput. +End Model0. + +Module Model. +Require typed_store_model. +(* This is the full model for Coqgen, including functions in the store *) +Module ModelTS := typed_store_model.ModelTypedStore (MLtypes). +(* Trying to build an instance with HB fails here too *) End Model. -Module IMonadTS := MonadTypedStore (MLtypes). -Import MLtypes. -Import IMonadTS. +(* Our proofs just use the interface *) +(* One can use either the "safe" interface based on Model0 *) +Module MLtypes' := Model0.ModelTS.MLtypes'. +(* or the standard one, allowing functions in the store *) +(* Module MLtypes' := Model.ModelTS.MLtypes'. *) + +Module IMonadTS := MonadTypedStore (MLtypes'). +Import MLtypes MLtypes' IMonadTS. Section cyclic. Variable M : typedStoreMonad. -Notation coq_type := (@MLtypes.coq_type M). +Notation coq_type := (@MLtypes'.coq_type M). Notation "'do' x <- m ; e" := (m >>= (fun x => e)) (at level 60, x name, m at level 200, e at level 60). Definition cycle (T : ml_type) (a b : coq_type T) : M (coq_type (ml_rlist T)) := - do r <- cnew (ml_rlist T) (Nil (coq_type T) T); + do r <- @cnew M (ml_rlist T) (Nil (coq_type T) T); do l <- (do v <- cnew (ml_rlist T) (Cons (coq_type T) T b r); Ret (Cons (coq_type T) T a v)); diff --git a/monad_model.v b/monad_model.v index 16373ead..214e3c93 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1719,18 +1719,17 @@ Module Type MLTYweak. Parameter ml_type : Set. Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. Parameter coq_type0 : ml_type -> Type. +Variant loc : ml_type -> Type := mkloc T : nat -> loc T. Parameter ml_undef : ml_type. -Parameter undef : coq_type0 ml_undef. +Parameter undef : (UU0 -> UU0) -> coq_type0 ml_undef. (*Parameter ml_exn : ml_type.*) End MLTYweak. Module ModelTypedStore (MLtypes : MLTYweak). Import MLtypes. Module MLtypes'. -Definition ml_type := ml_type. -Definition ml_type_eq_dec := ml_type_eq_dec. +Include MLtypes. Definition coq_type (M : UU0 -> UU0) := coq_type0. -Variant loc : ml_type -> Type := mkloc T : nat -> loc T. Definition locT := [eqType of nat]. Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. End MLtypes'. @@ -1751,7 +1750,7 @@ Local Notation M := acto. (*Local Notation coq_type := (coq_type M).*) Local Notation coq_type' := (coq_type idfun). -Definition def : binding idfun := mkbind ml_undef undef. +Definition def : binding idfun := mkbind ml_undef (undef idfun). Local Notation nth_error := List.nth_error. Definition cnew T (v : coq_type M T) : M (loc T) := @@ -1804,7 +1803,7 @@ Definition ml_type_eq_mixin := EqMixin ml_type_eqP. Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. (* Prove the laws *) -Let cnewget T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A) : +Definition cnewget T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. (*apply/boolp.funext => st/=. @@ -1822,7 +1821,7 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. Qed. -Let cnewput T (s t : coq_type M T) A (k : loc T -> M A) : +Definition cnewput T (s t : coq_type M T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. apply/boolp.funext => st /=. @@ -1836,7 +1835,7 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. Qed. -Let cgetput T (r : loc T) (s : coq_type M T) : +Definition cgetput T (r : loc T) (s : coq_type M T) : cget r >> cput r s = cput r s. Proof. apply/boolp.funext => st /=. @@ -1854,7 +1853,7 @@ case: ml_type_eq_dec => // HT. by rewrite -eq_rect_eq. Qed. -Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. +Definition cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. Proof. apply/boolp.funext => st /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. @@ -1868,7 +1867,8 @@ case: (ml_type_eq_dec T T) => H //. by rewrite -eq_rect_eq nth_error_set_nth_id. Qed. -Let cgetget T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A) : +Definition cgetget T (r : loc T) (A : UU0) + (k : coq_type M T -> coq_type M T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. apply/boolp.funext => st /=. @@ -1885,7 +1885,8 @@ case: (ml_type_eq_dec T T) => H //. by rewrite -eq_rect_eq. Qed. -Let cputget T (r: loc T) (s: coq_type M T) (A: UU0) (k: coq_type M T -> M A) : +Definition cputget T (r: loc T) (s: coq_type M T) (A: UU0) + (k: coq_type M T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. apply/boolp.funext => st /=. @@ -1904,7 +1905,7 @@ case: ml_type_eq_dec => H /=; last by rewrite bindfailf. by rewrite -eq_rect_eq !bindretf. Qed. -Let cputput T (r : loc T) (s s' : coq_type M T) : +Definition cputput T (r : loc T) (s s' : coq_type M T) : cput r s >> cput r s' = cput r s'. Proof. apply/boolp.funext => st. @@ -1920,7 +1921,7 @@ rewrite -eq_rect_eq. by rewrite set_set_nth eqxx. Qed. -Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) +Definition cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) (k : coq_type M T1 -> coq_type M T2 -> M A) : cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). @@ -1962,7 +1963,7 @@ case: (ml_type_eq_dec T1 T1) => // H. by rewrite -eq_rect_eq. Qed. -Let cgetnewD T T' (r : loc T) (s : coq_type M T') A +Definition cgetnewD T T' (r : loc T) (s : coq_type M T') A (k : loc T' -> coq_type M T -> coq_type M T -> M A) : cget r >>= (fun u => cnew s >>= fun r' => cget r >>= k r' u) = cget r >>= (fun u => cnew s >>= fun r' => k r' u u). @@ -1982,7 +1983,7 @@ case: ml_type_eq_dec => // H. by rewrite -eq_rect_eq. Qed. -Let cgetnewE T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) +Definition cgetnewE T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) (k1 k2 : loc T2 -> M A) : (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2). @@ -1999,7 +2000,7 @@ rewrite fmapE /= bindA /= 2!bindE /=. by rewrite Hk // neq_ltn /= (nth_error_size Hr1). Qed. -Let cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : +Definition cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip. Proof. apply/boolp.funext => st /=. @@ -2044,7 +2045,7 @@ rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. by case: ml_type_eq_dec. Qed. -Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) +Definition cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) (s2 : coq_type M T2) (A : UU0) : loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. @@ -2111,7 +2112,7 @@ case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq set_set_nth (negbTE Hr). Qed. -Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) +Definition cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) (A : UU0) (k : coq_type M T2 -> M A) : loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). @@ -2151,7 +2152,7 @@ case: ml_type_eq_dec => // HT1. by rewrite -eq_rect_eq bindE. Qed. -Let cputnewC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A +Definition cputnewC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A (k : loc T' -> M A) : cget r >> (cnew s' >>= fun r' => cput r s >> k r') = cput r s >> (cnew s' >>= k). @@ -2175,24 +2176,24 @@ rewrite (nth_error_size_set_nth _ _ Hr). by rewrite (nth_error_set_nth_rcons _ _ _ Hr). Qed. -Let crunret (A B : UU0) (m : M A) (s : B) : +Definition crunret (A B : UU0) (m : M A) (s : B) : crun m -> crun (m >> Ret s) = Some s. Proof. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. by case Hm: (m [::]). Qed. -Let crunskip : crun skip = Some tt. +Definition crunskip : crun skip = Some tt. Proof. by []. Qed. -Let crunnew (A : UU0) T (m : M A) (s : A -> coq_type M T) : +Definition crunnew (A : UU0) T (m : M A) (s : A -> coq_type M T) : crun m -> crun (m >>= fun x => cnew (s x)). Proof. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. by case Hm: (m [::]). Qed. -Let crunnewget (A : UU0) T (m : M A) (s : A -> coq_type M T) : +Definition crunnewget (A : UU0) T (m : M A) (s : A -> coq_type M T) : crun m -> crun (m >>= fun x => cnew (s x) >>= @cget T). Proof. rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. @@ -2200,7 +2201,7 @@ case Hm: (m [::]) => [|[a b]] // _. by rewrite bindE /= bindE /= -(bindmret (_ >>= _)) bindA cnewget bindE. Qed. -Let crungetnew (A : UU0) T1 T2 (m : M A) (r : A -> loc T1) +Definition crungetnew (A : UU0) T1 T2 (m : M A) (r : A -> loc T1) (s : A -> coq_type M T2) : crun (m >>= fun x => cget (r x)) -> crun (m >>= fun x => cnew (s x) >> cget (r x)). @@ -2214,7 +2215,8 @@ rewrite (nth_error_rcons_some _ Hnth). by case Hcoe: coerce. Qed. -Let crungetput (A : UU0) T (m : M A) (r : A -> loc T) (s : A -> coq_type M T) : +Definition crungetput (A : UU0) T (m : M A) (r : A -> loc T) + (s : A -> coq_type M T) : crun (m >>= fun x => cget (r x)) -> crun (m >>= fun x => cput (r x) (s x)). Proof. From 96fb1a778e17bb26b1b7d87390baeb7044e7134b Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 7 Jul 2023 12:23:12 +0900 Subject: [PATCH 72/82] rename MLtypes in a more coherent way --- example_typed_store.v | 18 +++++++++--------- monad_model.v | 2 +- typed_store_model.v | 2 +- 3 files changed, 11 insertions(+), 11 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 93bf2319..1d30c339 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -24,7 +24,7 @@ Inductive ml_type : Set := Inductive undef_t : Set := Undef. -Module MLtypes. +Module MLtypes0. Definition ml_type_eq_dec (T1 T2 : ml_type) : {T1=T2}+{T1<>T2}. revert T2; induction T1; destruct T2; try (right; intro; discriminate); try (now left); @@ -63,14 +63,14 @@ Fixpoint coq_type (T : ml_type) : Type := End with_monad. Local Definition coq_type0 := @coq_type idfun. Local Definition ml_undef := ml_undef. -End MLtypes. +End MLtypes0. (* Check that the models can indeed be instanciated, for soundness *) Module Model0. Require monad_model. (* This model does not allow functions in the store, but can be defined without endangering soundness. *) -Module ModelTS := monad_model.ModelTypedStore (MLtypes). +Module ModelTS := monad_model.ModelTypedStore (MLtypes0). (* Attempt to build an instance with HB *) Require Import fail_lib state_lib trace_lib. @@ -94,22 +94,22 @@ End Model0. Module Model. Require typed_store_model. (* This is the full model for Coqgen, including functions in the store *) -Module ModelTS := typed_store_model.ModelTypedStore (MLtypes). +Module ModelTS := typed_store_model.ModelTypedStore (MLtypes0). (* Trying to build an instance with HB fails here too *) End Model. (* Our proofs just use the interface *) (* One can use either the "safe" interface based on Model0 *) -Module MLtypes' := Model0.ModelTS.MLtypes'. +Module MLtypes := Model0.ModelTS.MLtypes'. (* or the standard one, allowing functions in the store *) -(* Module MLtypes' := Model.ModelTS.MLtypes'. *) +(* Module MLtypes := Model.ModelTS.MLtypes'. *) -Module IMonadTS := MonadTypedStore (MLtypes'). -Import MLtypes MLtypes' IMonadTS. +Module MonadTS := MonadTypedStore (MLtypes). +Import MLtypes0 MLtypes MonadTS. Section cyclic. Variable M : typedStoreMonad. -Notation coq_type := (@MLtypes'.coq_type M). +Notation coq_type := (@MLtypes.coq_type M). Notation "'do' x <- m ; e" := (m >>= (fun x => e)) (at level 60, x name, m at level 200, e at level 60). diff --git a/monad_model.v b/monad_model.v index 214e3c93..786814be 100644 --- a/monad_model.v +++ b/monad_model.v @@ -2230,7 +2230,7 @@ move/(_ (s a)). by case: coerce. Qed. -Canonical Structure isMonadTypedStoreModel := +Definition isMonadTypedStoreModel : isMonadTypedStore M := isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC cputgetC cputnewC diff --git a/typed_store_model.v b/typed_store_model.v index 94fdef3c..d680b72d 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -592,7 +592,7 @@ move/(_ (s a)). by case: coerce. Qed. -Definition isMonadTypedStoreModel := +Definition isMonadTypedStoreModel : isMonadTypedStore M := isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC cputgetC cputnewC From c3c900fc4a7f566dcc0d4b169b4fd937ed540780 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Fri, 14 Jul 2023 10:44:47 +0900 Subject: [PATCH 73/82] make loc contents abstract in hierarchy.v --- hierarchy.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/hierarchy.v b/hierarchy.v index ada8e05f..ef22c0ea 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1084,8 +1084,8 @@ HB.structure Definition MonadPlusArray (S : UU0) (I : eqType) := Module Type MLTY. Parameter ml_type : Set. Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Variant loc : ml_type -> Type := mkloc T : nat -> loc T. Parameter locT : eqType. +Variant loc : ml_type -> Type := mkloc T : locT -> loc T. Parameter loc_id : forall T, loc T -> locT. Parameter coq_type : forall M : Type -> Type, ml_type -> Type. End MLTY. From 81f9c777222331b05b10b8849af581255c525f8f Mon Sep 17 00:00:00 2001 From: affeldt-aist <33154536+affeldt-aist@users.noreply.github.com> Date: Fri, 14 Jul 2023 10:58:55 +0900 Subject: [PATCH 74/82] cgetputC (wip) + wip (#109) --- typed_store_model.v | 233 +++++++++++++++++++++----------------------- 1 file changed, 112 insertions(+), 121 deletions(-) diff --git a/typed_store_model.v b/typed_store_model.v index d680b72d..1e04c3df 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -49,7 +49,7 @@ Definition M0 Env (T : UU0) := MS Env option_monad T. #[bypass_check(positivity)] Inductive Env := mkEnv : seq (binding (M0 Env)) -> Env. Definition ofEnv '(mkEnv e) := e. -Definition sizeEnv '(mkEnv e) := size e. +Definition sizeEnv e := size (ofEnv e). Definition acto : UU0 -> UU0 := MS Env option_monad. Local Notation M := acto. @@ -83,7 +83,13 @@ Lemma coerce_eq (T T' : ml_type) (s : coq_type T) : coerce T' s -> T = T'. Proof. by rewrite /coerce; case: ml_type_eq_dec. Qed. Lemma coerce_Some (T : ml_type) (s : coq_type T) : coerce T s = Some s. -Proof. by rewrite /coerce; case: ml_type_eq_dec => // ?; rewrite -eq_rect_eq. Qed. +Proof. +by rewrite /coerce; case: ml_type_eq_dec => // ?; rewrite -eq_rect_eq. +Qed. + +Lemma coerce_None (T T' : ml_type) (s : coq_type T) : + T <> T' -> coerce T' s = None. +Proof. by rewrite /coerce; case: ml_type_eq_dec. Qed. Definition cget T (r : loc T) : M (coq_type T) := fun st => @@ -93,12 +99,11 @@ Definition cget T (r : loc T) : M (coq_type T) := Definition cput T (r : loc T) (v : coq_type T) : M unit := fun st => - let: mkEnv st := st in let n := loc_id r in - if nth_error st n is Some (mkbind T' _) then + if nth_error (ofEnv st) n is Some (mkbind T' _) then if coerce T' v is Some u then let b := mkbind (u : coq_type T') in - inr (tt, mkEnv (set_nth def st n b)) + inr (tt, mkEnv (set_nth def (ofEnv st) n b)) else inl tt else inl tt. @@ -149,8 +154,7 @@ Lemma cnewget T (s : coq_type T) A (k : loc T -> coq_type T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. apply/boolp.funext => e. -rewrite bind_cnew (Some_cget s)//. -by destruct e as [e]; rewrite nth_error_rcons_size. +by rewrite bind_cnew (Some_cget s)// nth_error_rcons_size. Qed. Lemma nocoerce_cput T (r : loc T) (s : coq_type T) T' (s' : coq_type T') e : @@ -171,9 +175,9 @@ Proof. by move=> H; move: e H => [e] H; rewrite /cput H. Qed. Lemma cnewput T (s t : coq_type T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. -apply/boolp.funext => -[st]. -rewrite bind_cnew 2!MS_bindE /= nth_error_rcons_size coerce_Some. -by rewrite set_nth_rcons. +apply/boolp.funext => e. +rewrite bind_cnew 2!MS_bindE. +by rewrite /cput/= nth_error_rcons_size coerce_Some set_nth_rcons. Qed. (* @@ -223,14 +227,19 @@ have [s' H|T' s' H Ts'|H] := ntherrorP e r. - by rewrite MS_bindE None_cget// None_cput. Qed. +(* TODO: at least rename *) +Lemma Some_cputE T (r : loc T) (s s' : coq_type T) e : + nth_error (ofEnv e) (loc_id r) = Some (mkbind s') -> + cput r s e = inr (tt, mkEnv (set_nth def (ofEnv e) (loc_id r) (mkbind s))). +Proof. by move=> H; rewrite /cput/= H coerce_Some. Qed. + Lemma Some_cput T (r : loc T) (s : coq_type T) e : nth_error (ofEnv e) (loc_id r) = Some (mkbind s) -> cput r s e = (@skip M) e. Proof. move=> H. -destruct e as [e]. -rewrite /cput/= H coerce_Some/=. -by rewrite nth_error_set_nth_id. +rewrite /cput/= H coerce_Some/= nth_error_set_nth_id//. +by destruct e. Qed. Lemma Some_cputget T (s' s : coq_type T) (r : loc T) A (k : coq_type T -> M A) @@ -239,13 +248,8 @@ Lemma Some_cputget T (s' s : coq_type T) (r : loc T) A (k : coq_type T -> M A) (cput r s >> (cget r >>= k)) e = (cput r s >> k s) e. Proof. move=> H. -destruct e as [e]. -rewrite {1}/cput/= !MS_bindE H. -rewrite coerce_Some. -rewrite bindE/=. -rewrite (Some_cget s); last first. - by rewrite /= nth_error_set_nth. -by rewrite H coerce_Some. +rewrite 2!MS_bindE (Some_cputE _ H). +by rewrite bindE/= (Some_cget s)// nth_error_set_nth. Qed. Arguments Some_cputget {T} s'. @@ -284,10 +288,10 @@ Lemma Some_cputput (T : ml_type) (r : loc T) (s s' : coq_type T) (cput r s >> cput r s') e = cput r s' e. Proof. move=> H. -destruct e as [e]. rewrite MS_bindE/=. -rewrite H !coerce_Some. -rewrite bindE/=. +rewrite [in LHS](Some_cputE _ H)//. +rewrite [in RHS](Some_cputE _ H)//. +rewrite bindE/= /cput. rewrite nth_error_set_nth coerce_Some/=. by rewrite set_set_nth eqxx. Qed. @@ -362,49 +366,46 @@ have [u r1u|T' s' Hr1 T1s'|Hr] := ntherrorP e r1. - by rewrite !MS_bindE None_cget. Qed. +(* TODO: do not rely on bindE to pass option and do not unfold cget *) Lemma cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type T2) : cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip. Proof. -apply/boolp.funext => -[st] /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= 2!bindA /= /cget /cput. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. - rewrite bindfailf. - case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; - last by rewrite bindfailf. - case Hc: (coerce T2' s) => [u|]; last by rewrite bindfailf. - rewrite bindE /= /bindS /= bindE /=. - by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). -case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. - rewrite bindfailf. - case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= Hr2. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec T1' T1) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec T2 T2') => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE/= bindE/=. - case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. - rewrite -Hr nth_error_set_nth /coerce. - move: Hr2; rewrite -Hr Hr1 => -[] HT2. - subst T1'. - by case: ml_type_eq_dec. - rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. - by case: ml_type_eq_dec. -rewrite bindE /= bindE /= Hr2 {1}/coerce. -case: ml_type_eq_dec => HT2 //. -subst T1' T2'. -rewrite -!eq_rect_eq. -rewrite bindE /= bindE /=. -case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. - rewrite -Hr nth_error_set_nth /coerce. - move: Hr2; rewrite -Hr Hr1 => -[] HT2. - subst T2. - by case: ml_type_eq_dec. -rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. -by case: ml_type_eq_dec. +apply/boolp.funext => e /=. +have [u r1u|T1' v1 Hr1 T1v1|Hr1] := ntherrorP e r1. +- rewrite (Some_cget u)// [in RHS]bindA MS_bindE. + have [v r2v|T' v r2v T2v|Hr2] := ntherrorP e r2. + + rewrite (Some_cputE _ r2v) [in RHS]bindE/=. + have [r12|r12] := eqVneq (loc_id r1) (loc_id r2). + move: r2v; rewrite -r12 r1u => -[?] _; subst T2. + by rewrite MS_bindE /cget nth_error_set_nth// coerce_Some. + by rewrite (Some_cget u)// (nth_error_set_nth_other _ _ r12 r1u). + + by rewrite (nocoerce_cput _ _ T2v). + + by rewrite None_cput. +- rewrite MS_bindE [RHS]MS_bindE bindA. + have [v r2v|T' v r2v T2v|Hr2] := ntherrorP e r2. + + rewrite {1}/cget Hr1. + have [?|T1'T1] := ml_type_eq_dec T1' T1. + * subst T1'; rewrite coerce_Some. + rewrite [in LHS]bindE/= (Some_cputE _ r2v) [in RHS]bindE/=. + have [Hr|Hr] := eqVneq (loc_id r1) (loc_id r2). + move: r2v; rewrite -Hr Hr1 => -[?] _; subst T2. + by rewrite /cget nth_error_set_nth coerce_Some. + by rewrite /cget (nth_error_set_nth_other _ _ Hr Hr1) coerce_Some. + * rewrite coerce_None// bindfailf (Some_cputE _ r2v) [in RHS]bindE/=. + have [r12|r12] := eqVneq (loc_id r1) (loc_id r2). + move: r2v; rewrite -r12 Hr1 => -[?] _; subst T1'. + by rewrite /cget nth_error_set_nth coerce_None. + by rewrite /cget (nth_error_set_nth_other _ _ r12 Hr1) coerce_None. + + by rewrite (nocoerce_cget _ T1v1)// bindfailf (nocoerce_cput _ _ T2v). + + by rewrite (nocoerce_cget _ T1v1)// None_cput// bindfailf. +- rewrite MS_bindE None_cget// bindfailf bindA MS_bindE. + have [v r2v|T' v r2v T2v|Hr2] := ntherrorP e r2. + + rewrite (Some_cputE _ r2v)/=. + rewrite bindE/=. + rewrite MS_bindE None_cget/= ?bindfailf//. + by rewrite (nth_error_set_nth_none _ _ Hr1 r2v). + + by rewrite (nocoerce_cput _ r2v). + + by rewrite None_cput. Qed. Lemma cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) @@ -413,65 +414,58 @@ Lemma cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. Proof. move=> H. -apply/boolp.funext => -[st] /=. +apply/boolp.funext => e /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. case Hr1: (nth_error _ _) => [[T1' u]|]; last first. rewrite bindfailf. case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. case Hc: (coerce T2' s2) => [u|]; last by rewrite bindfailf. - rewrite !bindE /=. - rewrite !bindE /=. + rewrite bindE /= bindE/=. case/boolP: (loc_id r1 == loc_id r2) => Hr. by rewrite -(eqP Hr) Hr1 in Hr2. by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). case/boolP: (loc_id r1 == loc_id r2) H => [/eqP|] Hr /= H. rewrite -Hr Hr1. case/boolP: (T1 == T2) => /eqP HT; last first. - rewrite /coerce. case: (ml_type_eq_dec T1 T1') => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec T2 T1') => HT2; last by rewrite bindfailf. + rewrite coerce_None// bindfailf. + case: (ml_type_eq_dec T2 T1') => ?; last by rewrite coerce_None. subst T1'. - rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. - by case: ml_type_eq_dec. + rewrite coerce_Some. + rewrite bindE/= bindE/=. + by rewrite nth_error_set_nth// coerce_None. subst T1'. - rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. - by case: ml_type_eq_dec => [/esym|] //. + rewrite coerce_Some bindE/= bindE/= nth_error_set_nth. + by rewrite coerce_None; auto. subst T2. - rewrite /coerce. - case: ml_type_eq_dec => HT1 //. + case: (ml_type_eq_dec T1 T1') => HT1 //; last first. + by rewrite coerce_None// bindE/= coerce_None. subst T1'. - rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /= !nth_error_set_nth. - case: ml_type_eq_dec => HT1 //. - rewrite -!eq_rect_eq. - case: H => // H. - by move: (JMeq_eq H) => <-. + rewrite coerce_Some bindE/= bindE/= bindE/= nth_error_set_nth. + rewrite coerce_Some/= bindE/= nth_error_set_nth coerce_Some. + by case: H => // H; rewrite -(JMeq_eq H). case Hr2: (nth_error _ _) => [[T2' v]|]; last first. rewrite bindfailf. case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. - rewrite !bindE /=. - rewrite !bindE /=. + rewrite bindE/= bindE/=. by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1). -rewrite {1 3}/coerce. -case: (ml_type_eq_dec _ _) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. +case: (ml_type_eq_dec T1 T1') => HT1; last first. + rewrite coerce_None// bindfailf. + case: (ml_type_eq_dec T2 T2') => HT2; last by rewrite coerce_None// bindfailf. subst T2'. - rewrite -eq_rect_eq. - rewrite bindE /=. - rewrite !bindE /=. - rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. - by case: ml_type_eq_dec. + rewrite coerce_Some /=. + rewrite bindE/= bindE/=. + by rewrite (nth_error_set_nth_other _ _ Hr Hr1) coerce_None. subst T1'. -rewrite -eq_rect_eq /coerce. -rewrite bindE /= bindE /= (nth_error_set_nth_other _ _ _ Hr2); - last by rewrite eq_sym. -case: ml_type_eq_dec => HT2' //. +rewrite coerce_Some. +rewrite bindE/= bindE/= (nth_error_set_nth_other _ _ _ Hr2) 1?eq_sym//. +case: (ml_type_eq_dec T2 T2') => HT2'; last by rewrite coerce_None. subst T2'. -rewrite -eq_rect_eq bindE /= bindE /= (nth_error_set_nth_other _ _ Hr Hr1). -case: ml_type_eq_dec => // HT1. -by rewrite -eq_rect_eq set_set_nth (negbTE Hr). +rewrite coerce_Some. +rewrite bindE/= bindE/= (nth_error_set_nth_other _ _ Hr Hr1). +rewrite coerce_Some. +by rewrite set_set_nth (negbTE Hr). Qed. Lemma cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) @@ -480,7 +474,7 @@ Lemma cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). Proof. move=> Hr. -apply/boolp.funext => -[st] /=. +apply/boolp.funext => e /=. rewrite !bindA. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. @@ -494,24 +488,23 @@ case Hr2: (nth_error _ _) => [[T2' v]|]; last first. case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE. by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1) bindfailf. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec _ _) => HT1; last first. +case: (ml_type_eq_dec T1 T1') => HT1; last first. + rewrite coerce_None//. rewrite bindfailf. - case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + case: (ml_type_eq_dec T2' T2) => HT2; last by rewrite coerce_None// bindfailf. subst T2'. - rewrite -eq_rect_eq. - rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE Hr1 /coerce. - by case: ml_type_eq_dec. + rewrite coerce_Some. + rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1. + by rewrite coerce_None. subst T1'. -rewrite -eq_rect_eq /coerce. +rewrite coerce_Some//. rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite (nth_error_set_nth_other _ _ _ Hr2); last by rewrite eq_sym. -case: ml_type_eq_dec => HT2' //. +rewrite (nth_error_set_nth_other _ _ _ Hr2) 1?eq_sym//. +case: (ml_type_eq_dec T2' T2) => HT2'; last by rewrite coerce_None. subst T2'. -rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /=. +rewrite coerce_Some bindE /= bindE /= bindE /= bindE /=. rewrite /bindS /= MS_mapE /= fmapE Hr1. -case: ml_type_eq_dec => // HT1. -by rewrite -eq_rect_eq bindE. +by rewrite coerce_Some. Qed. Lemma cputnewC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A @@ -519,22 +512,21 @@ Lemma cputnewC T T' (r : loc T) (s : coq_type T) (s' : coq_type T') A cget r >> (cnew s' >>= fun r' => cput r s >> k r') = cput r s >> (cnew s' >>= k). Proof. -apply/boolp.funext => -[st] /=. +apply/boolp.funext => e /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. +case: (ml_type_eq_dec T1 T) => H /=; last by rewrite coerce_None// coerce_None//; auto. subst T1. +rewrite coerce_Some//. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. rewrite (nth_error_rcons_some _ Hr). -rewrite /coerce. -case: ml_type_eq_dec => // H. -rewrite -eq_rect_eq. +rewrite coerce_Some. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. rewrite bindE /= /bindS /=. rewrite /fresh_loc/=. +rewrite /sizeEnv/=. rewrite (nth_error_size_set_nth _ _ Hr). by rewrite (nth_error_set_nth_rcons _ _ _ Hr). Qed. @@ -586,10 +578,9 @@ rewrite /crun /= !bindE /= /bindS !MS_mapE /= !fmapE /= !bindA /=. case Hm: (m _) => [|[a [b]]] //=. rewrite !bindE /= !bindE /= /cget /cput /=. case Hnth: nth_error => [[T' v]|] //. -case Hcoe: coerce => // _. -have /coerce_sym : is_true (coerce T v) by rewrite Hcoe. -move/(_ (s a)). -by case: coerce. +case: (ml_type_eq_dec T' T) => T'T; last by rewrite coerce_None. +subst T'. +by rewrite !coerce_Some. Qed. Definition isMonadTypedStoreModel : isMonadTypedStore M := From f913819faf366042c5cbc5db55e12b1bb1e5f796 Mon Sep 17 00:00:00 2001 From: affeldt-aist <33154536+affeldt-aist@users.noreply.github.com> Date: Mon, 24 Jul 2023 18:25:39 +0900 Subject: [PATCH 75/82] Typed store monad wip (#111) * cputC (wip) * refactoring * factorize coerce, uniformize monad_model/typed_store_model --- .github/workflows/docker-action.yml | 3 + README.md | 2 +- coq-monae.opam | 4 +- example_typed_store.v | 130 ++---- hierarchy.v | 91 ++-- meta.yml | 10 +- monad_model.v | 661 +++++++++++++++------------- typed_store_model.v | 231 +++++----- 8 files changed, 554 insertions(+), 578 deletions(-) diff --git a/.github/workflows/docker-action.yml b/.github/workflows/docker-action.yml index 30e53157..0d25bb46 100644 --- a/.github/workflows/docker-action.yml +++ b/.github/workflows/docker-action.yml @@ -17,7 +17,10 @@ jobs: strategy: matrix: image: + - 'mathcomp/mathcomp:1.15.0-coq-8.16' + - 'mathcomp/mathcomp:1.16.0-coq-8.16' - 'mathcomp/mathcomp:1.16.0-coq-8.17' + - 'mathcomp/mathcomp:1.17.0-coq-8.16' - 'mathcomp/mathcomp:1.17.0-coq-8.17' fail-fast: false steps: diff --git a/README.md b/README.md index 1e56bce3..826eaca2 100644 --- a/README.md +++ b/README.md @@ -26,7 +26,7 @@ in several examples of monadic equational reasoning. - Celestine Sauvage - Kazunari Tanaka - License: [LGPL-2.1-or-later](LICENSE) -- Compatible Coq versions: Coq 8.17 +- Compatible Coq versions: Coq 8.16-8.17 - Additional dependencies: - [MathComp ssreflect](https://math-comp.github.io) - [MathComp fingroup](https://math-comp.github.io) diff --git a/coq-monae.opam b/coq-monae.opam index 202e1a6d..3b86d8b9 100644 --- a/coq-monae.opam +++ b/coq-monae.opam @@ -15,10 +15,10 @@ description: """ This Coq library contains a hierarchy of monads with their laws used in several examples of monadic equational reasoning.""" -build: [make "-j%{jobs}%" ] +build: [make "-j%{jobs}%"] install: [make "install_full"] depends: [ - "coq" { (>= "8.17" & < "8.18~") | (= "dev") } + "coq" { (>= "8.16" & < "8.18~") | (= "dev") } "coq-mathcomp-ssreflect" { (>= "1.16.0" & < "1.18~") | (= "dev") } "coq-mathcomp-fingroup" { (>= "1.16.0" & < "1.18~") | (= "dev") } "coq-mathcomp-algebra" { (>= "1.16.0" & < "1.18~") | (= "dev") } diff --git a/example_typed_store.v b/example_typed_store.v index 1d30c339..0864e259 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -4,6 +4,7 @@ Require Import ZArith. From mathcomp Require Import all_ssreflect. From mathcomp Require boolp. From infotheo Require Import ssrZ. +Require monad_model. Require Import monae_lib hierarchy monad_lib. (*fail_lib state_lib.*) (******************************************************************************) @@ -13,18 +14,15 @@ Require Import monae_lib hierarchy monad_lib. (*fail_lib state_lib.*) Local Open Scope monae_scope. +Module MLTypes. Inductive ml_type : Set := | ml_int | ml_bool | ml_unit | ml_ref (_ : ml_type) | ml_arrow (_ : ml_type) (_ : ml_type) - | ml_rlist (_ : ml_type) - | ml_undef. - -Inductive undef_t : Set := Undef. + | ml_rlist (_ : ml_type). -Module MLtypes0. Definition ml_type_eq_dec (T1 T2 : ml_type) : {T1=T2}+{T1<>T2}. revert T2; induction T1; destruct T2; try (right; intro; discriminate); try (now left); @@ -36,86 +34,53 @@ revert T2; induction T1; destruct T2; right; injection; intros; contradiction. Defined. -Local Definition ml_type := ml_type. -Local Definition undef (M : UU0 -> UU0) := Undef. -Variant loc : ml_type -> Type := mkloc T : nat -> loc T. -Local Definition locT := [eqType of nat]. -Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. +Definition val_nonempty (M : UU0 -> UU0) := tt. + +Definition locT := [eqType of nat]. + +Notation loc := (@loc ml_type locT). Inductive rlist (a : Type) (a_1 : ml_type) := | Nil | Cons (_ : a) (_ : loc (ml_rlist a_1)). +End MLTypes. + +Module MLTypesNat. +Import MLTypes. + Section with_monad. Context [M : Type -> Type]. (* Generated type translation function *) -Fixpoint coq_type (T : ml_type) : Type := +Fixpoint coq_type_nat (T : ml_type) : Type := match T with | ml_int => nat | ml_bool => bool | ml_unit => unit - | ml_arrow T1 T2 => coq_type T1 -> M (coq_type T2) + | ml_arrow T1 T2 => coq_type_nat T1 -> M (coq_type_nat T2) | ml_ref T1 => loc T1 - | ml_rlist T1 => rlist (coq_type T1) T1 - | ml_undef => undef_t + | ml_rlist T1 => rlist (coq_type_nat T1) T1 end. End with_monad. -Local Definition coq_type0 := @coq_type idfun. -Local Definition ml_undef := ml_undef. -End MLtypes0. -(* Check that the models can indeed be instanciated, for soundness *) -Module Model0. -Require monad_model. -(* This model does not allow functions in the store, but can be defined - without endangering soundness. *) -Module ModelTS := monad_model.ModelTypedStore (MLtypes0). - -(* Attempt to build an instance with HB *) -Require Import fail_lib state_lib trace_lib. -Require Import monad_transformer monad_model. -Import ModelTS. -From HB Require Import structures. -(* This fails *) -Fail HB.instance Definition _ := - MTypedStore.isMonadTypedStore.Build _ cnewget cnewput cgetput cgetputskip - cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC - cputgetC cputnewC - crunret crunskip crunnew crunnewget crungetnew crungetput. -(* But we can still check that the definition is correct (again) *) -Definition isMonadTypedStoreModel := - MTypedStore.isMonadTypedStore.Build _ cnewget cnewput cgetput cgetputskip - cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC - cputgetC cputnewC - crunret crunskip crunnew crunnewget crungetnew crungetput. -End Model0. - -Module Model. -Require typed_store_model. -(* This is the full model for Coqgen, including functions in the store *) -Module ModelTS := typed_store_model.ModelTypedStore (MLtypes0). -(* Trying to build an instance with HB fails here too *) -End Model. - -(* Our proofs just use the interface *) -(* One can use either the "safe" interface based on Model0 *) -Module MLtypes := Model0.ModelTS.MLtypes'. -(* or the standard one, allowing functions in the store *) -(* Module MLtypes := Model.ModelTS.MLtypes'. *) - -Module MonadTS := MonadTypedStore (MLtypes). -Import MLtypes0 MLtypes MonadTS. +Definition ml_type_eq_mixin := comparableMixin MLTypes.ml_type_eq_dec. +Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. + +(*Canonical MLU := @Build_ML_universe _ coq_type ml_unit val_nonempty.*) + +Definition coq_type0 := @coq_type_nat idfun. +Canonical MLU := @monad_model.MLU _ coq_type0 ml_unit (val_nonempty idfun). Section cyclic. -Variable M : typedStoreMonad. -Notation coq_type := (@MLtypes.coq_type M). +Variables (M : typedStoreMonad MLU MLTypes.locT). +Notation coq_type := (@hierarchy.coq_type MLU M). Notation "'do' x <- m ; e" := (m >>= (fun x => e)) (at level 60, x name, m at level 200, e at level 60). Definition cycle (T : ml_type) (a b : coq_type T) : M (coq_type (ml_rlist T)) := - do r <- @cnew M (ml_rlist T) (Nil (coq_type T) T); + do r <- cnew (ml_rlist T) (Nil (coq_type T) T); do l <- (do v <- cnew (ml_rlist T) (Cons (coq_type T) T b r); Ret (Cons (coq_type T) T a v)); @@ -144,8 +109,8 @@ Qed. End cyclic. Section factorial. -Variable M : typedStoreMonad. -Notation coq_type := (@MLtypes.coq_type M). +Variable M : typedStoreMonad MLU MLTypes.locT. +Notation coq_type := (@coq_type M). Fixpoint fact_ref (r : loc ml_int) (n : nat) : M unit := if n is m.+1 then cget r >>= fun p => cput r (n * p) >> fact_ref r m @@ -209,8 +174,8 @@ End forloop. Arguments forloop {M}. Section fact_for. -Variable M : typedStoreMonad. -Notation coq_type := (@MLtypes.coq_type M). +Variable M : typedStoreMonad MLU MLTypes.locT. +Notation coq_type := (@coq_type MLU M). Notation "'do' x <- m ; e" := (m >>= (fun x => e)) (at level 60, x name, m at level 200, e at level 60). @@ -252,8 +217,8 @@ Qed. End fact_for. Section fibonacci. -Variable M : typedStoreMonad. -Notation coq_type := (@MLtypes.coq_type M). +Variable M : typedStoreMonad MLU MLTypes.locT. +Notation coq_type := (@coq_type M). Fixpoint fibo_rec n := if n is m.+1 then @@ -308,6 +273,8 @@ by under eq_bind do rewrite cnewput. Qed. End fibonacci. +End MLTypesNat. + Require Import PrimInt63. Require Sint63. @@ -458,38 +425,31 @@ Qed. End Int63. Module MLtypes63. -Local Definition ml_type_eq_dec := ml_type_eq_dec. -Local Definition ml_type := ml_type. -Local Definition undef (M : UU0 -> UU0) := Undef. -Variant loc : ml_type -> Set := mkloc T : nat -> loc T. -Local Definition locT := [eqType of nat]. -Local Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. +Import MLTypes. Section with_monad. Context [M : Type -> Type]. (* Generated type translation function *) -Fixpoint coq_type (T : ml_type) : Type := +Fixpoint coq_type63 (T : ml_type) : Type := match T with | ml_int => int | ml_bool => bool | ml_unit => unit - | ml_arrow T1 T2 => coq_type T1 -> M (coq_type T2) + | ml_arrow T1 T2 => coq_type63 T1 -> M (coq_type63 T2) | ml_ref T1 => loc T1 - | ml_rlist T1 => rlist (coq_type T1) T1 - | ml_undef => undef_t + | ml_rlist T1 => rlist (coq_type63 T1) T1 end. End with_monad. -Local Definition ml_undef := ml_undef. -End MLtypes63. -Module IMonadTS63 := hierarchy.MonadTypedStore (MLtypes63). -Import MLtypes63. -Import IMonadTS63. +Definition ml_type63_eq_mixin := comparableMixin MLTypes.ml_type_eq_dec. +Canonical ml_type63_eqType := Eval hnf in EqType _ ml_type63_eq_mixin. + +Canonical ml_type63 := @Build_ML_universe _ coq_type63 ml_unit val_nonempty. Section fact_for_int63. -Variable M : typedStoreMonad. -Notation coq_type := (@MLtypes63.coq_type M). +Variable M : typedStoreMonad ml_type63 MLTypes.locT. +Notation coq_type := (@MLtypes63.coq_type63 M). Notation "'do' x <- m ; e" := (m >>= (fun x => e)) (at level 60, x name, m at level 200, e at level 60). @@ -611,3 +571,5 @@ by rewrite lt0n subn_eq0 -ltnNge. Qed. End fact_for63_ok. End fact_for_int63. + +End MLtypes63. diff --git a/hierarchy.v b/hierarchy.v index ef22c0ea..88c49a14 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -1080,70 +1080,69 @@ HB.structure Definition MonadArray (S : UU0) (I : eqType) := HB.structure Definition MonadPlusArray (S : UU0) (I : eqType) := { M of MonadPlus M & isMonadArray S I M}. +Variant loc (ml_type : Type) (locT : eqType) : ml_type -> Type := + mkloc T : locT -> loc locT T. +Definition loc_id (ml_type : Type) (locT : eqType) {T : ml_type} (l : loc locT T) : locT := + let: mkloc _ n := l in n. + (* A monad for OCaml computations *) -Module Type MLTY. -Parameter ml_type : Set. -Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Parameter locT : eqType. -Variant loc : ml_type -> Type := mkloc T : locT -> loc T. -Parameter loc_id : forall T, loc T -> locT. -Parameter coq_type : forall M : Type -> Type, ml_type -> Type. -End MLTY. - -Module MonadTypedStore (MLtypes : MLTY). -Import MLtypes. - -HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) +Structure ML_universe := { + ml_type :> eqType ; + coq_type : forall M : Type -> Type, ml_type -> Type ; + ml_nonempty : ml_type ; + val_nonempty : forall M, coq_type M ml_nonempty }. + +HB.mixin Record isMonadTypedStore (MLU: ML_universe) (locT : eqType) (M : UU0 -> UU0) of Monad M := { - cnew : forall {T}, coq_type M T -> M (loc T) ; - cget : forall {T}, loc T -> M (coq_type M T) ; - cput : forall {T}, loc T -> coq_type M T -> M unit ; + cnew : forall {T : MLU}, coq_type M T -> M (loc locT T) ; + cget : forall {T}, loc locT T -> M (coq_type M T) ; + cput : forall {T}, loc locT T -> coq_type M T -> M unit ; crun : forall {A : UU0}, M A -> option A ; (* execute in empty store *) - cnewget : forall T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A), + cnewget : forall T (s : coq_type M T) A (k : loc locT T -> coq_type M T -> M A), cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s) ; - cnewput : forall T (s t : coq_type M T) A (k : loc T -> M A), + cnewput : forall T (s t : coq_type M T) A (k : loc locT T -> M A), cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k ; - cgetput : forall T (r : loc T) (s : coq_type M T), + cgetput : forall T (r : loc locT T) (s : coq_type M T), cget r >> cput r s = cput r s ; - cgetputskip : forall T (r : loc T), cget r >>= cput r = cget r >> skip ; + cgetputskip : forall T (r : loc locT T), cget r >>= cput r = cget r >> skip ; cgetget : - forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), + forall T (r : loc locT T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s ; cputget : - forall T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A), + forall T (r : loc locT T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A), cput r s >> (cget r >>= k) = cput r s >> k s ; - cputput : forall T (r : loc T) (s s' : coq_type M T), + cputput : forall T (r : loc locT T) (s s' : coq_type M T), cput r s >> cput r s' = cput r s' ; cgetC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) + forall T1 T2 (r1 : loc locT T1) (r2 : loc locT T2) (A : UU0) (k : coq_type M T1 -> coq_type M T2 -> M A), cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)) ; cgetnewD : - forall T T' (r : loc T) (s : coq_type M T') A - (k : loc T' -> coq_type M T -> coq_type M T -> M A), + forall T T' (r : loc locT T) (s : coq_type M T') A + (k : loc locT T' -> coq_type M T -> coq_type M T -> M A), cget r >>= (fun u => cnew s >>= (fun r' => cget r >>= k r' u)) = cget r >>= (fun u => cnew s >>= (fun r' => k r' u u)) ; - cgetnewE : forall T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) - (k1 k2 : loc T2 -> M A), - (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cgetnewE : forall T1 T2 (r1 : loc locT T1) (s : coq_type M T2) (A : UU0) + (k1 k2 : loc locT T2 -> M A), + (forall r2 : loc locT T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2) ; - cgetputC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2), + cgetputC : forall T1 T2 (r1 : loc locT T1) (r2 : loc locT T2) (s : coq_type M T2), cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip ; cputC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + forall T1 T2 (r1 : loc locT T1) (r2 : loc locT T2) (s1 : coq_type M T1) (s2 : coq_type M T2) (A : UU0), loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1 ; cputgetC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) + forall T1 T2 (r1 : loc locT T1) (r2 : loc locT T2) (s1 : coq_type M T1) (A : UU0) (k : coq_type M T2 -> M A), loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v) ; cputnewC : - forall T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A - (k : loc T' -> M A), + forall T T' (r : loc locT T) (s : coq_type M T) (s' : coq_type M T') A + (k : loc locT T' -> M A), cget r >> (cnew s' >>= fun r' => cput r s >> k r') = cput r s >> (cnew s' >>= k) ; crunret : forall (A B : UU0) (m : M A) (s : B), @@ -1154,27 +1153,28 @@ HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) crun m -> crun (m >>= fun x => cnew (s x)) ; crunnewget : forall (A : UU0) T (m : M A) (s : A -> coq_type M T), crun m -> crun (m >>= fun x => cnew (s x) >>= cget) ; - crungetnew : forall (A : UU0) T1 T2 (m : M A) (r : A -> loc T1) + crungetnew : forall (A : UU0) T1 T2 (m : M A) (r : A -> loc locT T1) (s : A -> coq_type M T2), crun (m >>= fun x => cget (r x)) -> crun (m >>= fun x => cnew (s x) >> cget (r x)) ; - crungetput : forall (A : UU0) T (m : M A) (r : A -> loc T) + crungetput : forall (A : UU0) T (m : M A) (r : A -> loc locT T) (s : A -> coq_type M T), crun (m >>= fun x => cget (r x)) -> crun (m >>= fun x => cput (r x) (s x)) ; }. #[short(type=typedStoreMonad)] -HB.structure Definition MonadTypedStore := - { M of isMonadTypedStore M & isMonad M & isFunctor M }. +HB.structure Definition MonadTypedStore (ml_type : ML_universe) (locT : eqType) := + { M of isMonadTypedStore ml_type locT M & isMonad M & isFunctor M }. -Arguments cnew {s}. -Arguments cget {s} [T]. -Arguments cput {s} [T]. -Arguments crun {s} [A]. +Arguments cnew {ml_type locT s}. +Arguments cget {ml_type locT s} [T]. +Arguments cput {ml_type locT s} [T]. +Arguments crun {ml_type locT s} [A]. Section cchk. -Variable M : typedStoreMonad. +Variables (MLU : ML_universe) (locT : eqType) (M : typedStoreMonad MLU locT). +Notation loc := (@loc MLU locT). Definition cchk T (r : loc T) : M unit := cget r >> skip. Lemma cnewchk T s (A : UU0) (k : loc T -> M A) : @@ -1251,11 +1251,10 @@ Proof. by rewrite bindmret. Qed. Lemma crunnew0 T s : crun (cnew T s : M _). Proof. by rewrite -(bindskipf (cnew T s)) crunnew // crunskip. Qed. -Lemma cnewgetret T s : cnew T s >>= @cget _ _ = cnew T s >> Ret s :> M _. +Lemma cnewgetret T s : cnew T s >>= @cget _ _ _ _ = cnew T s >> Ret s :> M _. Proof. under eq_bind do rewrite -cgetret; by rewrite cnewget. Qed. End cchk. -Arguments cchk {M} [T]. -End MonadTypedStore. +Arguments cchk {MLU locT M} [T]. HB.mixin Record isMonadTrace (T : UU0) (M : UU0 -> UU0) of Monad M := { mark : T -> M unit }. diff --git a/meta.yml b/meta.yml index 45c127e9..b343be44 100644 --- a/meta.yml +++ b/meta.yml @@ -35,12 +35,18 @@ license: file: LICENSE supported_coq_versions: - text: Coq 8.17 - opam: '{ (>= "8.17" & < "8.18~") | (= "dev") }' + text: Coq 8.16-8.17 + opam: '{ (>= "8.16" & < "8.18~") | (= "dev") }' tested_coq_opam_versions: +- version: '1.15.0-coq-8.16' + repo: 'mathcomp/mathcomp' +- version: '1.16.0-coq-8.16' + repo: 'mathcomp/mathcomp' - version: '1.16.0-coq-8.17' repo: 'mathcomp/mathcomp' +- version: '1.17.0-coq-8.16' + repo: 'mathcomp/mathcomp' - version: '1.17.0-coq-8.17' repo: 'mathcomp/mathcomp' diff --git a/monad_model.v b/monad_model.v index 786814be..23774811 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1715,61 +1715,69 @@ Arguments nth_error_size {T st n a}. Require Import JMeq. -Module Type MLTYweak. -Parameter ml_type : Set. -Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Parameter coq_type0 : ml_type -> Type. -Variant loc : ml_type -> Type := mkloc T : nat -> loc T. -Parameter ml_undef : ml_type. -Parameter undef : (UU0 -> UU0) -> coq_type0 ml_undef. -(*Parameter ml_exn : ml_type.*) -End MLTYweak. - -Module ModelTypedStore (MLtypes : MLTYweak). -Import MLtypes. -Module MLtypes'. -Include MLtypes. -Definition coq_type (M : UU0 -> UU0) := coq_type0. +(*Structure ML_universe0 := { + ml_type0 :> eqType ; + coq_type0 : ml_type0 -> Type ; + ml_nonempty0 : ml_type0 ; + val_nonempty0 : coq_type0 ml_nonempty0 }.*) + +Section coerce. +Variables (X : eqType) (f : X -> UU0). + +Definition coerce (T1 T2 : X) (v : f T1) : option (f T2) := + if @eqP _ T1 T2 is ReflectT H then Some (eq_rect _ _ v _ H) else None. + +Lemma coerce_sym (T T' : X) (s : f T) (s' : f T') : coerce T' s -> coerce T s'. +Proof. +by rewrite /coerce; case: eqP => //= h; case: eqP => //; rewrite h; auto. +Qed. + +Lemma coerce_Some (T : X) (s : f T) : coerce T s = Some s. +Proof. +by rewrite /coerce; case: eqP => /= [?|]; [rewrite -eq_rect_eq|auto]. +Qed. + +Lemma coerce_eq (T T' : X) (s : f T) : coerce T' s -> T = T'. +Proof. by rewrite /coerce; case: eqP. Qed. + +Lemma coerce_None (T T' : X) (s : f T) : T != T' -> coerce T' s = None. +Proof. by rewrite /coerce; case: eqP. Qed. + +End coerce. + +Section ModelTypedStore. +(*Variable MLU0 : ML_universe0.*) +Variables + (ml_type0 : eqType) + (coq_type0 : ml_type0 -> Type) + (ml_nonempty0 : ml_type0) + (val_nonempty0 : coq_type0 ml_nonempty0). + +Definition MLU : ML_universe := + @Build_ML_universe (ml_type0) (fun M => @coq_type0) + (ml_nonempty0) (fun M => val_nonempty0). + Definition locT := [eqType of nat]. -Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. -End MLtypes'. -Module MTypedStore := MonadTypedStore (MLtypes'). -Import MLtypes'. -Import MLtypes. -Import MTypedStore. Record binding (M : Type -> Type) := - mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. + mkbind { bind_type : MLU; bind_val : coq_type M bind_type }. Arguments mkbind {M}. Section typed_store. -Definition acto : UU0 -> UU0 := - (*MX (coq_type0 ml_exn)*) - (MS (seq (binding idfun)) [the monad of option_monad]). +Definition acto : UU0 -> UU0 := MS (seq (binding idfun)) [the monad of option_monad]. Local Notation M := acto. -(*Local Notation coq_type := (coq_type M).*) -Local Notation coq_type' := (coq_type idfun). +Local Notation coq_type' := (@coq_type MLU idfun). + +Definition def : binding idfun := mkbind (ml_nonempty MLU) (val_nonempty MLU idfun). -Definition def : binding idfun := mkbind ml_undef (undef idfun). Local Notation nth_error := List.nth_error. +Notation loc := (@loc MLU locT). + Definition cnew T (v : coq_type M T) : M (loc T) := fun st => let n := size st in Ret (mkloc T n, rcons st (mkbind T (v : coq_type' T))). -Definition coerce T1 T2 (v : coq_type M T1) : option (coq_type M T2) := - match ml_type_eq_dec T1 T2 with - | left H => Some (eq_rect _ _ v _ H) - | right _ => None - end. - -Lemma coerce_sym (T T' : ml_type) (s : coq_type M T) (s' : coq_type M T') : - coerce T' s -> coerce T s'. -Proof. -rewrite /coerce. -case: ml_type_eq_dec => [h|h]; last by case: ml_type_eq_dec. -by case: ml_type_eq_dec => [//|g]; subst T'. -Qed. Definition cget T (r : loc T) : M (coq_type M T) := fun st => @@ -1792,15 +1800,27 @@ Definition crun (A : UU0) (m : M A) : option A := | inr (a, _) => Some a end. -(* Make ml_type an eqType *) -Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. -Lemma ml_type_eqP : Equality.axiom ml_type_eqb. -Proof. -rewrite /ml_type_eqb => T1 T2. -by case: ml_type_eq_dec; constructor. -Qed. -Definition ml_type_eq_mixin := EqMixin ml_type_eqP. -Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. +Definition Env := seq (binding idfun). + +Local Notation mkbind := (@mkbind idfun). + +Definition extend_env T (v : coq_type M T) (e : Env) := + rcons e (mkbind T v). +Definition fresh_loc (T : MLU) (e : Env) := mkloc T (size e). + +Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M B) s : + (m >>= f) s = m s >>= uncurry f. +Proof. by []. Qed. + +Lemma bind_cnew T (s : coq_type M T) A (k : loc T -> M A) e : + (cnew s >>= k) e = k (fresh_loc T e) (extend_env s e). +Proof. by case: e. Qed. + +Lemma Some_cget T (r : loc T) (s : coq_type M T) e (A : UU0) (f : coq_type M T -> M A) : + nth_error e (loc_id r) = Some (mkbind T s) -> + (cget r >>= f) e = f s e. +Proof. by move=> H; rewrite MS_bindE /cget H coerce_Some. Qed. +Arguments Some_cget {T r} s. (* Prove the laws *) Definition cnewget T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A) : @@ -1813,112 +1833,146 @@ transitivity ( ((stateT (seq (binding idfun)) option_monad # (fun r : loc T => cget r >>= k r)) (cnew s)) st ) => //.*) -apply/boolp.funext => st/=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite bindE /= bindE /= bindE /= MS_mapE /= /bindS /= fmapE /=. -rewrite /cget nth_error_rcons_size /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq. +apply/boolp.funext => e. +by rewrite bind_cnew (Some_cget s)// nth_error_rcons_size. Qed. Definition cnewput T (s t : coq_type M T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE bindA. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= [in RHS]bindE /= /cput. -rewrite nth_error_rcons_size. -rewrite /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. +apply/boolp.funext => e. +rewrite bind_cnew 2!MS_bindE. +by rewrite /cput/= nth_error_rcons_size coerce_Some set_nth_rcons. +Qed. + +Variant nth_error_spec (T : MLU) (e : Env) (r : loc T) : Type := + | NthError : forall s : coq_type M T, + nth_error e (loc_id r) = Some (mkbind T s) -> nth_error_spec e r + | NthError_nocoerce : forall T' (s' : coq_type M T'), + nth_error e (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> + nth_error_spec e r + | NthError_None : nth_error e (loc_id r) = None -> nth_error_spec e r. + +Lemma ntherrorP (T : MLU) (e : Env) (r : loc T) : nth_error_spec e r. +Proof. +case H : (nth_error e (loc_id r)) => [[T' s']|]. + have [Ts'|Ts'] := boolp.pselect (coerce T s'). + have ? := coerce_eq Ts'; subst T'. + exact: NthError H. + exact: NthError_nocoerce H Ts'. +exact: NthError_None. Qed. +Lemma nocoerce_cget T (r : loc T) T' (s' : coq_type M T') e : + nth_error e (loc_id r) = Some (mkbind T' s') -> + ~ coerce T s' -> + cget r e = fail. +Proof. by move=> H Ts'; rewrite /cget H; case: coerce Ts'. Qed. + +Lemma nocoerce_cput T (r : loc T) (s : coq_type M T) T' (s' : coq_type M T') e : + nth_error e (loc_id r) = Some (mkbind T' s') -> + ~ coerce T s' -> + cput r s e = fail. +Proof. +move=> H Ts'; rewrite /cput H. +have {}Ts' : ~ coerce T' s by apply: contra_not Ts'; exact: coerce_sym. +by case: coerce Ts'. +Qed. + +Lemma None_cget T (r : loc T) e : + nth_error e (loc_id r) = None -> cget r e = fail. +Proof. by move=> H; rewrite /cget H. Qed. + +Lemma None_cput T (r : loc T) (s : coq_type M T) e : + nth_error e (loc_id r) = None -> cput r s e = fail. +Proof. by move=> H; rewrite /cput H. Qed. + Definition cgetput T (r : loc T) (s : coq_type M T) : cget r >> cput r s = cput r s. Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -rewrite /coerce. -case: ml_type_eq_dec => HT; last by case: ml_type_eq_dec => // /esym //. -subst T'. -rewrite -eq_rect_eq. -case: ml_type_eq_dec => // HT. -rewrite -eq_rect_eq {HT}. -do! rewrite bindE /=. -rewrite Hr. -case: ml_type_eq_dec => // HT. -by rewrite -eq_rect_eq. +apply/boolp.funext => e. +have [s' H|T' s' H Ts'|H] := ntherrorP e r. +- by rewrite (Some_cget s'). +- by rewrite MS_bindE (nocoerce_cget H)// (nocoerce_cput _ H). +- by rewrite MS_bindE None_cget// None_cput. +Qed. + +Lemma Some_cput T (r : loc T) (s : coq_type M T) e : + nth_error e (loc_id r) = Some (mkbind T s) -> + cput r s e = (@skip M) e. +Proof. +by move=> H; rewrite /cput/= H coerce_Some/= nth_error_set_nth_id. Qed. Definition cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE bindA /=. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite /coerce. -case: (ml_type_eq_dec T1 T) => H //. -subst T1. -rewrite bindE /= bindE /= Hr. -case: (ml_type_eq_dec T T) => H //. -by rewrite -eq_rect_eq nth_error_set_nth_id. +apply/boolp.funext => e /=. +have [s' H|T' s' H Ts'|H] := ntherrorP e r. +- by rewrite (Some_cget s')// (Some_cget s')// (Some_cput H). +- by rewrite !MS_bindE (nocoerce_cget H). +- by rewrite !MS_bindE None_cget. Qed. Definition cgetget T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite /coerce. -case: (ml_type_eq_dec T1 T) => H //. -subst T1. -rewrite bindE /= bindE /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite Hr. -case: (ml_type_eq_dec T T) => H //. -by rewrite -eq_rect_eq. +apply/boolp.funext => e /=. +have [s' H|T' s' H Ts'|H] := ntherrorP e r. +- by do 3 rewrite (Some_cget s')//. +- by rewrite !MS_bindE (nocoerce_cget H). +- by rewrite !MS_bindE None_cget. +Qed. + +Lemma Some_cputE T (r : loc T) (s s' : coq_type M T) e : + nth_error e (loc_id r) = Some (mkbind T s') -> + cput r s e = inr (tt, set_nth def e (loc_id r) (mkbind T s)). +Proof. by move=> H; rewrite /cput/= H coerce_Some. Qed. + +Lemma Some_cputget T (s' s : coq_type M T) (r : loc T) A (k : coq_type M T -> M A) + (e : Env) : + nth_error e (loc_id r) = Some (mkbind T s') -> + (cput r s >> (cget r >>= k)) e = (cput r s >> k s) e. +Proof. +move=> H. +rewrite 2!MS_bindE (Some_cputE _ H). +by rewrite bindE/= (Some_cget s)// nth_error_set_nth. Qed. +Arguments Some_cputget {T} s'. Definition cputget T (r: loc T) (s: coq_type M T) (A: UU0) (k: coq_type M T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. -apply/boolp.funext => st /=. -case: r s k => {}T n s k /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. -case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T'. -rewrite !bindretf /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite nth_error_set_nth. -rewrite /coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -by rewrite -eq_rect_eq !bindretf. +apply/boolp.funext => e /=. +have [s' H|T' s' H Ts'|H] := ntherrorP e r. +- by rewrite (Some_cputget s'). +- by rewrite !MS_bindE (nocoerce_cput _ H). +- by rewrite 2!MS_bindE None_cput. +Qed. + +Lemma Some_cputput (T : MLU) (r : loc T) (s s' : coq_type M T) + (e : Env) (s'' : coq_type M T) : + nth_error e (loc_id r) = Some (mkbind T s'') -> + (cput r s >> cput r s') e = cput r s' e. +Proof. +move=> H. +rewrite MS_bindE/=. +rewrite [in LHS](Some_cputE _ H)//. +rewrite [in RHS](Some_cputE _ H)//. +rewrite bindE/= /cput. +rewrite nth_error_set_nth coerce_Some/=. +by rewrite set_set_nth eqxx. Qed. Definition cputput T (r : loc T) (s s' : coq_type M T) : cput r s >> cput r s' = cput r s'. Proof. -apply/boolp.funext => st. -case: r s s' => {}T n s s' /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. -case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T'. -rewrite !bindretf /= nth_error_set_nth /coerce. -case: ml_type_eq_dec => // H. -rewrite -eq_rect_eq. -by rewrite set_set_nth eqxx. +apply/boolp.funext => e. +have [s'' H|T'' s'' H Ts''|H] := ntherrorP e r. +- by rewrite (Some_cputput _ _ H). +- by rewrite !MS_bindE (nocoerce_cput _ H)// (nocoerce_cput _ H). +- by rewrite MS_bindE !None_cput. Qed. Definition cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) @@ -1926,41 +1980,33 @@ Definition cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr1: (nth_error _ _) => [[T1' u]|]; last first. - rewrite bindfailf. - case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. - case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. - rewrite !bindE /=. - by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. -case Hr2: (nth_error _ _) => [[T2' v]|]; last first. - rewrite bindfailf. - case Hc: (coerce T1 u) => [v|]; last by rewrite bindfailf. - rewrite !bindE /=. - by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr2 bindfailf. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec _ _) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE /=. - rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr1 {1}/coerce. - case: (ml_type_eq_dec _ _) => //; by rewrite bindfailf. -subst T1'. -rewrite -eq_rect_eq. -rewrite bindE /=. -rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr2 {1}/coerce. -case: (ml_type_eq_dec T2' T2) => HT2; last by rewrite !bindfailf. -subst T2'. -rewrite -eq_rect_eq. -rewrite !bindE /=. -rewrite !bindE /=. -rewrite /bindS MS_mapE /= fmapE /= bindA /= Hr1 /coerce. -case: (ml_type_eq_dec T1 T1) => // H. -by rewrite -eq_rect_eq. +apply/boolp.funext => e /=. +have [u Hr1|u T1' Hr1 T1u|Hr1] := ntherrorP e r1. +- rewrite (Some_cget u)//. + have [v Hr2|v T2' Hr2 T2v|Hr2] := ntherrorP e r2. + + by rewrite (Some_cget v)// (Some_cget v)// (Some_cget u). + + by rewrite 2!MS_bindE (nocoerce_cget Hr2)// bindfailf. + + by rewrite !MS_bindE None_cget. +- rewrite MS_bindE (nocoerce_cget Hr1)// bindfailf. + have [v Hr2|v T2' Hr2 T2v|Hr2] := ntherrorP e r2. + + by rewrite (Some_cget v)// MS_bindE (nocoerce_cget Hr1). + + by rewrite MS_bindE (nocoerce_cget Hr2). + + by rewrite MS_bindE None_cget. +- rewrite MS_bindE None_cget// bindfailf. + have [v Hr2|v T2' Hr2 T2v|Hr2] := ntherrorP e r2. + + by rewrite (Some_cget v)// MS_bindE None_cget. + + by rewrite MS_bindE (nocoerce_cget Hr2). + + by rewrite MS_bindE !None_cget. +Qed. + +(* NB: this is similar to the cnewget law *) +Lemma cnewgetD_helper e T T' r v (s : coq_type M T') A (k : loc T' -> coq_type M T -> M A) : + nth_error e (loc_id r) = Some (mkbind T v) -> + (cnew s >>= (fun r' => cget r >>= k r')) e = (cnew s >>= (fun r => k r v)) e. +Proof. +move=> H. +rewrite bind_cnew//. +by rewrite (Some_cget v) // (nth_error_rcons_some _ H). Qed. Definition cgetnewD T T' (r : loc T) (s : coq_type M T') A @@ -1968,19 +2014,11 @@ Definition cgetnewD T T' (r : loc T) (s : coq_type M T') A cget r >>= (fun u => cnew s >>= fun r' => cget r >>= k r' u) = cget r >>= (fun u => cnew s >>= fun r' => k r' u u). Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T1. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite (nth_error_rcons_some _ Hr). -rewrite /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq. +apply/boolp.funext => e. +have [u Hr|u T1 Hr T1u|Hr] := ntherrorP e r. +- by rewrite (Some_cget u)// (Some_cget u)// (cnewgetD_helper _ _ Hr)//. +- by rewrite !MS_bindE (nocoerce_cget Hr). +- by rewrite !MS_bindE None_cget. Qed. Definition cgetnewE T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) @@ -1989,60 +2027,77 @@ Definition cgetnewE T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2). Proof. move=> Hk. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] //=. -case Hc: (coerce T1 v1) => [u1|] //. -rewrite bindE /= bindE /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= [in RHS]bindE /= [in RHS]bindE /= /bindS MS_mapE /=. -rewrite fmapE /= bindA /= 2!bindE /=. -by rewrite Hk // neq_ltn /= (nth_error_size Hr1). +apply/boolp.funext => e. +have [u r1u|T' s' Hr1 T1s'|Hr] := ntherrorP e r1. +- rewrite (Some_cget u)// (Some_cget u)// !bind_cnew Hk// neq_ltn. +- by move: r1u => /= /nth_error_size ->. +- by rewrite !MS_bindE (nocoerce_cget Hr1). +- by rewrite !MS_bindE None_cget. Qed. Definition cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip. Proof. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= 2!bindA /= /cget /cput. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. - rewrite bindfailf. - case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; - last by rewrite bindfailf. - case Hc: (coerce T2' s) => [u|]; last by rewrite bindfailf. - rewrite bindE /= /bindS /= bindE /=. - by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). -case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. - rewrite bindfailf. - case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= Hr2. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec T1' T1) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec T2 T2') => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE/= bindE/=. - case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. - rewrite -Hr nth_error_set_nth /coerce. - move: Hr2; rewrite -Hr Hr1 => -[] HT2. - subst T1'. - by case: ml_type_eq_dec. - rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. - by case: ml_type_eq_dec. -rewrite bindE /= bindE /= Hr2 {1}/coerce. -case: ml_type_eq_dec => HT2 //. -subst T1' T2'. -rewrite -!eq_rect_eq. -rewrite bindE /= bindE /=. -case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. - rewrite -Hr nth_error_set_nth /coerce. - move: Hr2; rewrite -Hr Hr1 => -[] HT2. - subst T2. - by case: ml_type_eq_dec. -rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. -by case: ml_type_eq_dec. +apply/boolp.funext => e /=. +have [u r1u|T1' v1 Hr1 T1v1|Hr1] := ntherrorP e r1. +- rewrite (Some_cget u)// [in RHS]bindA MS_bindE. + have [v r2v|T' v r2v T2v|Hr2] := ntherrorP e r2. + + rewrite (Some_cputE _ r2v) [in RHS]bindE/=. + have [r12|r12] := eqVneq (loc_id r1) (loc_id r2). + move: r2v; rewrite -r12 r1u => -[?] _; subst T2. + by rewrite MS_bindE /cget nth_error_set_nth// coerce_Some. + by rewrite (Some_cget u)// (nth_error_set_nth_other _ _ r12 r1u). + + by rewrite (nocoerce_cput _ _ T2v). + + by rewrite None_cput. +- rewrite MS_bindE [RHS]MS_bindE bindA. + have [v r2v|T' v r2v T2v|Hr2] := ntherrorP e r2. + + rewrite {1}/cget Hr1. + have [?|T1'T1] := eqVneq T1' T1. + * subst T1'; rewrite coerce_Some. + rewrite [in LHS]bindE/= (Some_cputE _ r2v) [in RHS]bindE/=. + have [Hr|Hr] := eqVneq (loc_id r1) (loc_id r2). + move: r2v; rewrite -Hr Hr1 => -[?] _; subst T2. + by rewrite /cget nth_error_set_nth coerce_Some. + by rewrite /cget (nth_error_set_nth_other _ _ Hr Hr1) coerce_Some. + * rewrite coerce_None//= bindfailf (Some_cputE _ r2v) [in RHS]bindE/=. + have [r12|r12] := eqVneq (loc_id r1) (loc_id r2). + move: r2v; rewrite -r12 Hr1 => -[?] _; subst T1'. + by rewrite /cget nth_error_set_nth coerce_None. + by rewrite /cget (nth_error_set_nth_other _ _ r12 Hr1) coerce_None. + + by rewrite (nocoerce_cget _ T1v1)// bindfailf (nocoerce_cput _ _ T2v). + + by rewrite (nocoerce_cget _ T1v1)// None_cput// bindfailf. +- rewrite MS_bindE None_cget// bindfailf bindA MS_bindE. + have [v r2v|T' v r2v T2v|Hr2] := ntherrorP e r2. + + rewrite (Some_cputE _ r2v)/=. + rewrite bindE/=. + rewrite MS_bindE None_cget/= ?bindfailf//. + by rewrite (nth_error_set_nth_none _ _ Hr1 r2v). + + by rewrite (nocoerce_cput _ r2v). + + by rewrite None_cput. +Qed. + +Lemma cput_then_fail T1 T2 T1' e (s1' : coq_type M T1') + (r2 : loc T2) (s2 : coq_type M T2) + (r1 : loc T1) (s1 : coq_type M T1) : + nth_error e (loc_id r1) = Some (mkbind T1' s1') -> + ~ coerce T1 s1' -> + (cput r2 s2 >> cput r1 s1) e = fail. +Proof. +move=> Hr1 T1s'. +have [v Hr2|T2' v Hr2 T2v|Hr2] := ntherrorP e r2; last first. +- by rewrite MS_bindE None_cput. +- by rewrite MS_bindE (nocoerce_cput _ _ T2v). +- rewrite MS_bindE (Some_cputE _ Hr2). + rewrite bindE/=. + have [Hr|Hr] := eqVneq (loc_id r1) (loc_id r2). + + rewrite {1}/cput -Hr /= nth_error_set_nth. + have [HT|HT] := eqVneq T1 T2; last by rewrite coerce_None. + subst T2. + rewrite coerce_Some//=. + move: Hr1; rewrite Hr Hr2 => -[?]; subst T1' => _. + by rewrite coerce_Some // in T1s'. + + rewrite (nocoerce_cput _ _ T1s')//=. + by rewrite (@nth_error_set_nth_other _ _ _ _ _ (mkbind T1' s1')). Qed. Definition cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) @@ -2050,66 +2105,48 @@ Definition cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. Proof. -move=> H. -apply/boolp.funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. -case Hr1: (nth_error _ _) => [[T1' u]|]; last first. - rewrite bindfailf. - case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. - case Hc: (coerce T2' s2) => [u|]; last by rewrite bindfailf. - rewrite !bindE /=. - rewrite !bindE /=. - case/boolP: (loc_id r1 == loc_id r2) => Hr. - by rewrite -(eqP Hr) Hr1 in Hr2. - by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). -case/boolP: (loc_id r1 == loc_id r2) H => [/eqP|] Hr /= H. - rewrite -Hr Hr1. - case/boolP: (T1 == T2) => /eqP HT; last first. - rewrite /coerce. - case: (ml_type_eq_dec T1 T1') => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec T2 T1') => HT2; last by rewrite bindfailf. - subst T1'. - rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. - by case: ml_type_eq_dec. - subst T1'. - rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. - by case: ml_type_eq_dec => [/esym|] //. - subst T2. - rewrite /coerce. - case: ml_type_eq_dec => HT1 //. - subst T1'. - rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /= !nth_error_set_nth. - case: ml_type_eq_dec => HT1 //. - rewrite -!eq_rect_eq. - case: H => // H. - by move: (JMeq_eq H) => <-. -case Hr2: (nth_error _ _) => [[T2' v]|]; last first. - rewrite bindfailf. - case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. - rewrite !bindE /=. - rewrite !bindE /=. - by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1). -rewrite {1 3}/coerce. -case: (ml_type_eq_dec _ _) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE /=. - rewrite !bindE /=. - rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. - by case: ml_type_eq_dec. -subst T1'. -rewrite -eq_rect_eq /coerce. -rewrite bindE /= bindE /= (nth_error_set_nth_other _ _ _ Hr2); - last by rewrite eq_sym. -case: ml_type_eq_dec => HT2' //. -subst T2'. -rewrite -eq_rect_eq bindE /= bindE /= (nth_error_set_nth_other _ _ Hr Hr1). -case: ml_type_eq_dec => // HT1. -by rewrite -eq_rect_eq set_set_nth (negbTE Hr). +move=> H; apply/boolp.funext => e /=. +have [u Hr1|T1' s'd Hr1 T1s'|Hr1] := ntherrorP e r1; last first. + rewrite MS_bindE None_cput// bindfailf. + have [v Hr2|T2' s' Hr2 T2s'|Hr2] := ntherrorP e r2; last first. + by rewrite MS_bindE None_cput. + + by rewrite MS_bindE (nocoerce_cput _ _ T2s'). + + rewrite MS_bindE (Some_cputE _ Hr2). + rewrite bindE/=. + by rewrite None_cput// (nth_error_set_nth_none _ _ Hr1 Hr2). +- rewrite MS_bindE (nocoerce_cput _ _ T1s')// bindfailf. + by rewrite (cput_then_fail _ _ _ Hr1). +- rewrite MS_bindE (Some_cputE _ Hr1)/=. + rewrite bindE/=. + case/boolP: (loc_id r1 == loc_id r2) H => [/eqP|] Hr /=. + + case => // H. + rewrite MS_bindE. + rewrite {2}/cput -Hr Hr1. + case/boolP: (T1 == T2) => [/eqP HT|HT]; last first. + rewrite coerce_None//; last by rewrite eq_sym. + by rewrite /cput/= Hr nth_error_set_nth// coerce_None// eq_sym. + subst T2. + rewrite coerce_Some bindE/=. + have ? := JMeq_eq H; subst s2. + by rewrite /cput -Hr. + + move=> _. + have [v r2v|T2' s Hr2 T2s|Hr2] := ntherrorP e r2; last first. + * rewrite MS_bindE None_cput/=; last first. + by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1). + by rewrite None_cput. + * rewrite MS_bindE. + rewrite [in RHS](nocoerce_cput _ _ T2s)// bindfailf. + rewrite (nocoerce_cput _ _ T2s)//=. + by rewrite (nth_error_set_nth_other _ _ _ Hr2) 1?eq_sym. + * rewrite MS_bindE. + rewrite [in RHS](Some_cputE _ r2v). + rewrite {1}/cput/=. + rewrite (nth_error_set_nth_other _ _ _ r2v) 1?eq_sym//. + rewrite coerce_Some. + rewrite bindE/=. + rewrite /cput (nth_error_set_nth_other _ _ Hr Hr1). + rewrite coerce_Some. + by rewrite set_set_nth (negbTE Hr). Qed. Definition cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) @@ -2118,7 +2155,7 @@ Definition cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). Proof. move=> Hr. -apply/boolp.funext => st /=. +apply/boolp.funext => e /=. rewrite !bindA. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. @@ -2132,24 +2169,23 @@ case Hr2: (nth_error _ _) => [[T2' v]|]; last first. case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE. by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1) bindfailf. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec _ _) => HT1; last first. +case: (eqVneq T1 T1') => HT1; last first. + rewrite coerce_None//. rewrite bindfailf. - case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. + case: (eqVneq T2' T2) => HT2; last by rewrite coerce_None// bindfailf. subst T2'. - rewrite -eq_rect_eq. - rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE Hr1 /coerce. - by case: ml_type_eq_dec. + rewrite coerce_Some. + rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1. + by rewrite coerce_None. subst T1'. -rewrite -eq_rect_eq /coerce. +rewrite coerce_Some//. rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite (nth_error_set_nth_other _ _ _ Hr2); last by rewrite eq_sym. -case: ml_type_eq_dec => HT2' //. +rewrite (nth_error_set_nth_other _ _ _ Hr2) 1?eq_sym//. +case: (eqVneq T2' T2) => HT2'; last by rewrite coerce_None. subst T2'. -rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /=. +rewrite coerce_Some bindE /= bindE /= bindE /= bindE /=. rewrite /bindS /= MS_mapE /= fmapE Hr1. -case: ml_type_eq_dec => // HT1. -by rewrite -eq_rect_eq bindE. +by rewrite coerce_Some. Qed. Definition cputnewC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A @@ -2157,21 +2193,20 @@ Definition cputnewC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A cget r >> (cnew s' >>= fun r' => cput r s >> k r') = cput r s >> (cnew s' >>= k). Proof. -apply/boolp.funext => st /=. +apply/boolp.funext => e /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. +case: (eqVneq T1 T) => H //=; last by rewrite coerce_None// 1?eq_sym// coerce_None// eq_sym. subst T1. +rewrite [in LHS]coerce_Some//. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. rewrite (nth_error_rcons_some _ Hr). -rewrite /coerce. -case: ml_type_eq_dec => // H. -rewrite -eq_rect_eq. +rewrite coerce_Some. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. rewrite bindE /= /bindS /=. +rewrite /fresh_loc/=. rewrite (nth_error_size_set_nth _ _ Hr). by rewrite (nth_error_set_nth_rcons _ _ _ Hr). Qed. @@ -2221,28 +2256,22 @@ Definition crungetput (A : UU0) T (m : M A) (r : A -> loc T) crun (m >>= fun x => cput (r x) (s x)). Proof. rewrite /crun /= !bindE /= /bindS !MS_mapE /= !fmapE /= !bindA /=. -case Hm: (m [::]) => [|[a b]] //. -rewrite !bindE /= !bindE /= /cget /cput. +case Hm: (m _) => [|[a b]] //=. +rewrite !bindE /= !bindE /= /cget /cput /=. case Hnth: nth_error => [[T' v]|] //. -case Hcoe: coerce => // _. -have /coerce_sym : is_true (coerce T v) by rewrite Hcoe. -move/(_ (s a)). -by case: coerce. +case: (eqVneq T' T) => T'T; last first. + by rewrite coerce_None// 1?eq_sym// coerce_None. +subst T'. +by rewrite !coerce_Some. Qed. -Definition isMonadTypedStoreModel : isMonadTypedStore M := - isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip +HB.instance Definition _ := Monad.on M. +HB.instance Definition isMonadTypedStoreModel := + isMonadTypedStore.Build MLU locT M cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC cputgetC cputnewC crunret crunskip crunnew crunnewget crungetnew crungetput. -(* The elpi tactic/command HB.instance failed without giving a specific - error message. -HB.instance Definition _ := - isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip - cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC - cputgetC cputnewC crunret crunskip crunnew. -*) End typed_store. End ModelTypedStore. diff --git a/typed_store_model.v b/typed_store_model.v index 1e04c3df..901ee0d1 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -19,78 +19,49 @@ Unset Printing Implicit Defensive. Local Open Scope monae_scope. -Module Type MLTYdef. -Parameter ml_type : Set. -Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Variant loc : ml_type -> Type := mkloc T : nat -> loc T. -Parameter coq_type : forall M : Type -> Type, ml_type -> Type. -Parameter ml_undef : ml_type. -Parameter undef : forall M, coq_type M ml_undef. -End MLTYdef. - -Module ModelTypedStore (MLtypes : MLTYdef). -Import MLtypes. -Module MLtypes'. -Include MLtypes. -Definition locT := [eqType of nat]. -Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. -End MLtypes'. -Module MTypedStore := MonadTypedStore (MLtypes'). -Import MLtypes'. -Import MLtypes. -Import MTypedStore. +Section ModelTypedStore. +Variable MLU : ML_universe. Record binding (M : Type -> Type) := - mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. + mkbind { bind_type : MLU; bind_val : coq_type M bind_type }. Arguments mkbind {M bind_type}. Definition M0 Env (T : UU0) := MS Env option_monad T. +End ModelTypedStore. + #[bypass_check(positivity)] -Inductive Env := mkEnv : seq (binding (M0 Env)) -> Env. -Definition ofEnv '(mkEnv e) := e. +Inductive Env (MLU : ML_universe) := mkEnv : seq (binding MLU (M0 (Env _))) -> Env _. + +Section ModelTypedStore_contd. +Variable MLU : ML_universe. + +Definition ofEnv (e : Env MLU) := let: mkEnv e := e in e. Definition sizeEnv e := size (ofEnv e). +Local Notation mkEnv := (@mkEnv MLU). +Local Notation Env := (@Env MLU). + Definition acto : UU0 -> UU0 := MS Env option_monad. Local Notation M := acto. -Local Notation coq_type := (coq_type M). + +Local Notation coq_type := (@coq_type MLU M). + +Local Notation undef := (val_nonempty MLU). Definition def := mkbind (undef M). + Local Notation nth_error := List.nth_error. Definition extend_env T (v : coq_type T) (e : Env) := mkEnv (rcons (ofEnv e) (mkbind v)). -Definition fresh_loc (T : ml_type) (e : Env) := mkloc T (sizeEnv e). +Definition fresh_loc (T : MLU) (e : Env) := mkloc T (sizeEnv e). + +Local Notation loc := (@loc MLU locT). Definition cnew T (v : coq_type T) : M (loc T) := fun e => inr (fresh_loc T e, extend_env v e). -Definition coerce T1 T2 (v : coq_type T1) : option (coq_type T2) := - match ml_type_eq_dec T1 T2 with - | left H => Some (eq_rect _ _ v _ H) - | right _ => None - end. - -Lemma coerce_sym (T T' : ml_type) (s : coq_type T) (s' : coq_type T') : - coerce T' s -> coerce T s'. -Proof. -rewrite /coerce. -case: ml_type_eq_dec => [h|h]; last by case: ml_type_eq_dec. -by case: ml_type_eq_dec => [//|g]; subst T'. -Qed. - -Lemma coerce_eq (T T' : ml_type) (s : coq_type T) : coerce T' s -> T = T'. -Proof. by rewrite /coerce; case: ml_type_eq_dec. Qed. - -Lemma coerce_Some (T : ml_type) (s : coq_type T) : coerce T s = Some s. -Proof. -by rewrite /coerce; case: ml_type_eq_dec => // ?; rewrite -eq_rect_eq. -Qed. - -Lemma coerce_None (T T' : ml_type) (s : coq_type T) : - T <> T' -> coerce T' s = None. -Proof. by rewrite /coerce; case: ml_type_eq_dec. Qed. - Definition cget T (r : loc T) : M (coq_type T) := fun st => if nth_error (ofEnv st) (loc_id r) is Some (mkbind T' v) then @@ -113,16 +84,6 @@ Definition crun (A : UU0) (m : M A) : option A := | inr (a, _) => Some a end. -(* Make ml_type an eqType *) -Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. -Lemma ml_type_eqP : Equality.axiom ml_type_eqb. -Proof. -rewrite /ml_type_eqb => T1 T2. -by case: ml_type_eq_dec; constructor. -Qed. -Definition ml_type_eq_mixin := EqMixin ml_type_eqP. -Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. - (* WIP *) Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M B) s : (m >>= f) s = m s >>= uncurry f. @@ -200,7 +161,7 @@ rewrite !bind_cnew /fresh_loc /extend_env /= size_rcons. Abort. *) -Variant nth_error_spec (T : ml_type) (e : Env) (r : loc T) : Type := +Variant nth_error_spec (T : MLU) (e : Env) (r : loc T) : Type := | NthError : forall s : coq_type T, nth_error (ofEnv e) (loc_id r) = Some (mkbind s) -> nth_error_spec e r | NthError_nocoerce : forall T' (s' : coq_type T'), @@ -208,7 +169,7 @@ Variant nth_error_spec (T : ml_type) (e : Env) (r : loc T) : Type := nth_error_spec e r | NthError_None : nth_error (ofEnv e) (loc_id r) = None -> nth_error_spec e r. -Lemma ntherrorP (T : ml_type) (e : Env) (r : loc T) : nth_error_spec e r. +Lemma ntherrorP (T : MLU) (e : Env) (r : loc T) : nth_error_spec e r. Proof. case H : (nth_error (ofEnv e) (loc_id r)) => [[T' s']|]. have [Ts'|Ts'] := boolp.pselect (coerce T s'). @@ -282,7 +243,7 @@ have [s' H|T' s' H Ts'|H] := ntherrorP e r. - by rewrite 2!MS_bindE None_cput. Qed. -Lemma Some_cputput (T : ml_type) (r : loc T) (s s' : coq_type T) +Lemma Some_cputput (T : MLU) (r : loc T) (s s' : coq_type T) (e : Env) (s'' : coq_type T) : nth_error (ofEnv e) (loc_id r) = Some (mkbind s'') -> (cput r s >> cput r s') e = cput r s' e. @@ -384,14 +345,14 @@ have [u r1u|T1' v1 Hr1 T1v1|Hr1] := ntherrorP e r1. - rewrite MS_bindE [RHS]MS_bindE bindA. have [v r2v|T' v r2v T2v|Hr2] := ntherrorP e r2. + rewrite {1}/cget Hr1. - have [?|T1'T1] := ml_type_eq_dec T1' T1. + have [?|T1'T1] := eqVneq T1' T1. * subst T1'; rewrite coerce_Some. rewrite [in LHS]bindE/= (Some_cputE _ r2v) [in RHS]bindE/=. have [Hr|Hr] := eqVneq (loc_id r1) (loc_id r2). move: r2v; rewrite -Hr Hr1 => -[?] _; subst T2. by rewrite /cget nth_error_set_nth coerce_Some. by rewrite /cget (nth_error_set_nth_other _ _ Hr Hr1) coerce_Some. - * rewrite coerce_None// bindfailf (Some_cputE _ r2v) [in RHS]bindE/=. + * rewrite coerce_None//= bindfailf (Some_cputE _ r2v) [in RHS]bindE/=. have [r12|r12] := eqVneq (loc_id r1) (loc_id r2). move: r2v; rewrite -r12 Hr1 => -[?] _; subst T1'. by rewrite /cget nth_error_set_nth coerce_None. @@ -408,64 +369,77 @@ have [u r1u|T1' v1 Hr1 T1v1|Hr1] := ntherrorP e r1. + by rewrite None_cput. Qed. +Lemma cput_then_fail T1 T2 T1' e (s1' : coq_type T1') + (r2 : loc T2) (s2 : coq_type T2) + (r1 : loc T1) (s1 : coq_type T1) : + nth_error (ofEnv e) (loc_id r1) = Some (mkbind s1') -> + ~ coerce T1 s1' -> + (cput r2 s2 >> cput r1 s1) e = fail. +Proof. +move=> Hr1 T1s'. +have [v Hr2|T2' v Hr2 T2v|Hr2] := ntherrorP e r2; last first. +- by rewrite MS_bindE None_cput. +- by rewrite MS_bindE (nocoerce_cput _ _ T2v). +- rewrite MS_bindE (Some_cputE _ Hr2). + rewrite bindE/=. + have [Hr|Hr] := eqVneq (loc_id r1) (loc_id r2). + + rewrite {1}/cput -Hr /= nth_error_set_nth. + have [HT|HT] := eqVneq T1 T2; last by rewrite coerce_None. + subst T2. + rewrite coerce_Some//=. + move: Hr1; rewrite Hr Hr2 => -[?]; subst T1' => _. + by rewrite coerce_Some // in T1s'. + + rewrite (nocoerce_cput _ _ T1s')//=. + by rewrite (@nth_error_set_nth_other _ _ _ _ _ (mkbind s1')). +Qed. + Lemma cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) (s2 : coq_type T2) (A : UU0) : loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. Proof. -move=> H. -apply/boolp.funext => e /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. -case Hr1: (nth_error _ _) => [[T1' u]|]; last first. - rewrite bindfailf. - case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. - case Hc: (coerce T2' s2) => [u|]; last by rewrite bindfailf. - rewrite bindE /= bindE/=. - case/boolP: (loc_id r1 == loc_id r2) => Hr. - by rewrite -(eqP Hr) Hr1 in Hr2. - by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). -case/boolP: (loc_id r1 == loc_id r2) H => [/eqP|] Hr /= H. - rewrite -Hr Hr1. - case/boolP: (T1 == T2) => /eqP HT; last first. - case: (ml_type_eq_dec T1 T1') => HT1; last first. - rewrite coerce_None// bindfailf. - case: (ml_type_eq_dec T2 T1') => ?; last by rewrite coerce_None. - subst T1'. +move=> H; apply/boolp.funext => e /=. +have [u Hr1|T1' s'd Hr1 T1s'|Hr1] := ntherrorP e r1; last first. + rewrite MS_bindE None_cput// bindfailf. + have [v Hr2|T2' s' Hr2 T2s'|Hr2] := ntherrorP e r2; last first. + by rewrite MS_bindE None_cput. + + by rewrite MS_bindE (nocoerce_cput _ _ T2s'). + + rewrite MS_bindE (Some_cputE _ Hr2). + rewrite bindE/=. + by rewrite None_cput// (nth_error_set_nth_none _ _ Hr1 Hr2). +- rewrite MS_bindE (nocoerce_cput _ _ T1s')// bindfailf. + by rewrite (cput_then_fail _ _ _ Hr1). +- rewrite MS_bindE (Some_cputE _ Hr1)/=. + rewrite bindE/=. + case/boolP: (loc_id r1 == loc_id r2) H => [/eqP|] Hr /=. + + case => // H. + rewrite MS_bindE. + rewrite {2}/cput -Hr Hr1. + case/boolP: (T1 == T2) => [/eqP HT|HT]; last first. + rewrite coerce_None//; last by rewrite eq_sym. + by rewrite /cput/= Hr nth_error_set_nth// coerce_None// eq_sym. + subst T2. + rewrite coerce_Some bindE/=. + have ? := JMeq_eq H; subst s2. + by rewrite /cput -Hr. + + move=> _. + have [v r2v|T2' s Hr2 T2s|Hr2] := ntherrorP e r2; last first. + * rewrite MS_bindE None_cput/=; last first. + by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1). + by rewrite None_cput. + * rewrite MS_bindE. + rewrite [in RHS](nocoerce_cput _ _ T2s)// bindfailf. + rewrite (nocoerce_cput _ _ T2s)//=. + by rewrite (nth_error_set_nth_other _ _ _ Hr2) 1?eq_sym. + * rewrite MS_bindE. + rewrite [in RHS](Some_cputE _ r2v). + rewrite {1}/cput/=. + rewrite (nth_error_set_nth_other _ _ _ r2v) 1?eq_sym//. rewrite coerce_Some. - rewrite bindE/= bindE/=. - by rewrite nth_error_set_nth// coerce_None. - subst T1'. - rewrite coerce_Some bindE/= bindE/= nth_error_set_nth. - by rewrite coerce_None; auto. - subst T2. - case: (ml_type_eq_dec T1 T1') => HT1 //; last first. - by rewrite coerce_None// bindE/= coerce_None. - subst T1'. - rewrite coerce_Some bindE/= bindE/= bindE/= nth_error_set_nth. - rewrite coerce_Some/= bindE/= nth_error_set_nth coerce_Some. - by case: H => // H; rewrite -(JMeq_eq H). -case Hr2: (nth_error _ _) => [[T2' v]|]; last first. - rewrite bindfailf. - case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. - rewrite bindE/= bindE/=. - by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1). -case: (ml_type_eq_dec T1 T1') => HT1; last first. - rewrite coerce_None// bindfailf. - case: (ml_type_eq_dec T2 T2') => HT2; last by rewrite coerce_None// bindfailf. - subst T2'. - rewrite coerce_Some /=. - rewrite bindE/= bindE/=. - by rewrite (nth_error_set_nth_other _ _ Hr Hr1) coerce_None. -subst T1'. -rewrite coerce_Some. -rewrite bindE/= bindE/= (nth_error_set_nth_other _ _ _ Hr2) 1?eq_sym//. -case: (ml_type_eq_dec T2 T2') => HT2'; last by rewrite coerce_None. -subst T2'. -rewrite coerce_Some. -rewrite bindE/= bindE/= (nth_error_set_nth_other _ _ Hr Hr1). -rewrite coerce_Some. -by rewrite set_set_nth (negbTE Hr). + rewrite bindE/=. + rewrite /cput (nth_error_set_nth_other _ _ Hr Hr1). + rewrite coerce_Some. + by rewrite set_set_nth (negbTE Hr). Qed. Lemma cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type T1) @@ -488,10 +462,10 @@ case Hr2: (nth_error _ _) => [[T2' v]|]; last first. case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE. by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1) bindfailf. -case: (ml_type_eq_dec T1 T1') => HT1; last first. +case: (eqVneq T1 T1') => HT1; last first. rewrite coerce_None//. rewrite bindfailf. - case: (ml_type_eq_dec T2' T2) => HT2; last by rewrite coerce_None// bindfailf. + case: (eqVneq T2' T2) => HT2; last by rewrite coerce_None// bindfailf. subst T2'. rewrite coerce_Some. rewrite bindE/= bindE/= bindE/= /bindS /= MS_mapE /= fmapE Hr1. @@ -500,7 +474,7 @@ subst T1'. rewrite coerce_Some//. rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. rewrite (nth_error_set_nth_other _ _ _ Hr2) 1?eq_sym//. -case: (ml_type_eq_dec T2' T2) => HT2'; last by rewrite coerce_None. +case: (eqVneq T2' T2) => HT2'; last by rewrite coerce_None. subst T2'. rewrite coerce_Some bindE /= bindE /= bindE /= bindE /=. rewrite /bindS /= MS_mapE /= fmapE Hr1. @@ -516,9 +490,9 @@ apply/boolp.funext => e /=. rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -case: (ml_type_eq_dec T1 T) => H /=; last by rewrite coerce_None// coerce_None//; auto. +case: (eqVneq T1 T) => H //=; last by rewrite coerce_None// 1?eq_sym// coerce_None// eq_sym. subst T1. -rewrite coerce_Some//. +rewrite [in LHS]coerce_Some//. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. rewrite (nth_error_rcons_some _ Hr). @@ -578,14 +552,17 @@ rewrite /crun /= !bindE /= /bindS !MS_mapE /= !fmapE /= !bindA /=. case Hm: (m _) => [|[a [b]]] //=. rewrite !bindE /= !bindE /= /cget /cput /=. case Hnth: nth_error => [[T' v]|] //. -case: (ml_type_eq_dec T' T) => T'T; last by rewrite coerce_None. +case: (eqVneq T' T) => T'T; last first. + by rewrite coerce_None// 1?eq_sym// coerce_None. subst T'. by rewrite !coerce_Some. Qed. -Definition isMonadTypedStoreModel : isMonadTypedStore M := - isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip +HB.instance Definition _ := Monad.on M. +HB.instance Definition isMonadTypedStoreModel := + isMonadTypedStore.Build _ _ M cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC cputgetC cputnewC crunret crunskip crunnew crunnewget crungetnew crungetput. -End ModelTypedStore. + +End ModelTypedStore_contd. From 5fd0198be6e41cd4177364eaeb834d90c72f79cc Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Mon, 24 Jul 2023 19:17:53 +0900 Subject: [PATCH 76/82] fix --- .github/workflows/docker-action.yml | 1 - meta.yml | 2 -- 2 files changed, 3 deletions(-) diff --git a/.github/workflows/docker-action.yml b/.github/workflows/docker-action.yml index 0d25bb46..07d18eb4 100644 --- a/.github/workflows/docker-action.yml +++ b/.github/workflows/docker-action.yml @@ -17,7 +17,6 @@ jobs: strategy: matrix: image: - - 'mathcomp/mathcomp:1.15.0-coq-8.16' - 'mathcomp/mathcomp:1.16.0-coq-8.16' - 'mathcomp/mathcomp:1.16.0-coq-8.17' - 'mathcomp/mathcomp:1.17.0-coq-8.16' diff --git a/meta.yml b/meta.yml index b343be44..c30903f4 100644 --- a/meta.yml +++ b/meta.yml @@ -39,8 +39,6 @@ supported_coq_versions: opam: '{ (>= "8.16" & < "8.18~") | (= "dev") }' tested_coq_opam_versions: -- version: '1.15.0-coq-8.16' - repo: 'mathcomp/mathcomp' - version: '1.16.0-coq-8.16' repo: 'mathcomp/mathcomp' - version: '1.16.0-coq-8.17' From cdda5dadef1eb53b77790d7b1e5b364644aa35ff Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 25 Jul 2023 11:07:34 +0900 Subject: [PATCH 77/82] doc --- _CoqProject | 1 + example_typed_store.v | 17 +- hierarchy.v | 96 +----- impredicative_set/ihierarchy.v | 182 +---------- impredicative_set/imonad_lib.v | 10 + impredicative_set/imonad_model.v | 536 ------------------------------- monad_lib.v | 10 +- typed_store_lib.v | 99 ++++++ 8 files changed, 149 insertions(+), 802 deletions(-) create mode 100644 typed_store_lib.v diff --git a/_CoqProject b/_CoqProject index 752a4c65..a541dec3 100644 --- a/_CoqProject +++ b/_CoqProject @@ -13,6 +13,7 @@ state_lib.v array_lib.v trace_lib.v proba_lib.v +typed_store_lib.v monad_composition.v monad_model.v proba_monad_model.v diff --git a/example_typed_store.v b/example_typed_store.v index 0864e259..956c9591 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -5,11 +5,21 @@ From mathcomp Require Import all_ssreflect. From mathcomp Require boolp. From infotheo Require Import ssrZ. Require monad_model. -Require Import monae_lib hierarchy monad_lib. (*fail_lib state_lib.*) +Require Import monae_lib hierarchy monad_lib typed_store_lib. (******************************************************************************) -(* Type store examples *) +(* Typed store examples *) (* *) +(* Inductive ml_type *) +(* *) +(* Module MLTypesNat *) +(* Definition cycle *) +(* Fixpoint fact_ref *) +(* Definition fact_for *) +(* Fixpoint fibo_ref *) +(* *) +(* Module MLtypes63 *) +(* Definition fact_for63 *) (******************************************************************************) Local Open Scope monae_scope. @@ -537,8 +547,7 @@ rewrite /N2int -!Sint63.is_int //. by split; apply N2int_bounded, Z.lt_le_incl, (Z.lt_trans _ _ _ Hm). Qed. -Theorem fact_for63_ok : - crun (fact_for63 (N2int n)) = Some (N2int (fact_rec n)). +Theorem fact_for63_ok : crun (fact_for63 (N2int n)) = Some (N2int (fact_rec n)). Proof. rewrite /fact_for63. under eq_bind do rewrite !bindA !bindretf. diff --git a/hierarchy.v b/hierarchy.v index 88c49a14..4da4c461 100644 --- a/hierarchy.v +++ b/hierarchy.v @@ -65,6 +65,11 @@ From HB Require Import structures. (* arrayMonad == array monad *) (* plusArrayMonad == plus monad + array monad *) (* *) +(* ML_universe == a type with decidable equality to represent an *) +(* OCaml type together with its Coq representation *) +(* in the type of a Tarski universe *) +(* typedStoreMonad == A monad for OCaml computations *) +(* *) (* Trace monads: *) (* traceMonad == trace monad *) (* traceReifyMonad == trace + reify *) @@ -83,8 +88,10 @@ From HB Require Import structures. (* failFreshMonad == freshMonad + failure *) (* *) (* references: *) +(* - R. Affeldt, J. Garrigue, T. Saikawa, Environment-friendly monadic *) +(* equational reasoning for OCaml, The Coq Workshop 2023 *) (* - R. Affeldt, D. Nowak, Extending Equational Monadic Reasoning with Monad *) -(* Transformers, https://arxiv.org/abs/2011.03463 *) +(* Transformers, TYPES 2020, https://arxiv.org/abs/2011.03463 *) (* - R. Affeldt, D. Nowak, T. Saikawa, A Hierarchy of Monadic Effects for *) (* Program Verification using Equational Reasoning, MPC 2019 *) (* - J. Gibbons, R. Hinze, Just do it: simple monadic equational reasoning, *) @@ -1085,14 +1092,13 @@ Variant loc (ml_type : Type) (locT : eqType) : ml_type -> Type := Definition loc_id (ml_type : Type) (locT : eqType) {T : ml_type} (l : loc locT T) : locT := let: mkloc _ n := l in n. -(* A monad for OCaml computations *) Structure ML_universe := { ml_type :> eqType ; coq_type : forall M : Type -> Type, ml_type -> Type ; ml_nonempty : ml_type ; val_nonempty : forall M, coq_type M ml_nonempty }. -HB.mixin Record isMonadTypedStore (MLU: ML_universe) (locT : eqType) (M : UU0 -> UU0) +HB.mixin Record isMonadTypedStore (MLU : ML_universe) (locT : eqType) (M : UU0 -> UU0) of Monad M := { cnew : forall {T : MLU}, coq_type M T -> M (loc locT T) ; cget : forall {T}, loc locT T -> M (coq_type M T) ; @@ -1172,90 +1178,6 @@ Arguments cget {ml_type locT s} [T]. Arguments cput {ml_type locT s} [T]. Arguments crun {ml_type locT s} [A]. -Section cchk. -Variables (MLU : ML_universe) (locT : eqType) (M : typedStoreMonad MLU locT). -Notation loc := (@loc MLU locT). -Definition cchk T (r : loc T) : M unit := cget r >> skip. - -Lemma cnewchk T s (A : UU0) (k : loc T -> M A) : - cnew T s >>= (fun r => cchk r >> k r) = cnew T s >>= k. -Proof. -under eq_bind do rewrite bindA. -rewrite cnewget. -by under eq_bind do rewrite bindskipf. -Qed. - -Lemma cchknewC T1 T2 (r : loc T1) s (A : UU0) (k : loc T2 -> M A) : - cchk r >> (cnew T2 s >>= fun r' => cchk r >> k r') = - cchk r >> (cnew T2 s >>= k). -Proof. -rewrite !(bindA,bindskipf). -under eq_bind do under eq_bind do rewrite bindA. -rewrite cgetnewD. -by under eq_bind do under eq_bind do rewrite bindskipf. -Qed. - -Lemma cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) - (k: coq_type M T2 -> M A) : - cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). -Proof. -under [in RHS]eq_bind do rewrite bindA bindskipf. -by rewrite !(bindA,bindskipf) cgetC. -Qed. - -Lemma cchknewE T1 T2 (r1 : loc T1) s (A : UU0) (k1 k2 : loc T2 -> M A) : - (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> - cchk r1 >> (cnew T2 s >>= k1) = cchk r1 >> (cnew T2 s >>= k2). -Proof. move=> Hk; rewrite !(bindA,bindskipf); exact: cgetnewE. Qed. - -Lemma cchknewget T T' (r : loc T) s (A : UU0) k : - cchk r >> (cnew T' s >>= fun r' => cget r >>= k r') = - cget r >>= (fun u => cnew T' s >>= k ^~ u) :> M A. -Proof. by rewrite bindA bindskipf cgetnewD. Qed. - -Lemma cchknewput T T' (r : loc T) s s' (A : UU0) k : - cchk r >> (cnew T' s' >>= fun r' => cput r s >> k r') = - cput r s >> (cnew T' s' >>= k) :> M A. -Proof. by rewrite bindA bindskipf cputnewC. Qed. - -Lemma cchkget T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : - cchk r >> (cget r >>= k) = cget r >>= k. -Proof. by rewrite bindA bindskipf cgetget. Qed. - -Lemma cgetchk T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : - cget r >>= (fun s => cchk r >> k s) = cget r >>= k. -Proof. under eq_bind do rewrite bindA bindskipf; by rewrite cgetget. Qed. - -Lemma cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : - cchk r1 >> cput r2 s = cput r2 s >> cchk r1. -Proof. by rewrite bindA bindskipf cgetputC bindA. Qed. - -Lemma cchkput T (r : loc T) (s : coq_type M T) : - cchk r >> cput r s = cput r s. -Proof. by rewrite bindA bindskipf cgetput. Qed. - -Lemma cputchk T (r : loc T) (s : coq_type M T) : - cput r s >> cchk r = cput r s. -Proof. by rewrite cputget bindmskip. Qed. - -Lemma cchkC T1 T2 (r1: loc T1) (r2: loc T2) : - cchk r1 >> cchk r2 = cchk r2 >> cchk r1. -Proof. by rewrite !(bindA,bindskipf) cgetC. Qed. - -Lemma cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. -Proof. by rewrite bindA bindskipf cgetget. Qed. - -Lemma cgetret T (r : loc T) : cget r >>= Ret = cget r :> M _. -Proof. by rewrite bindmret. Qed. - -Lemma crunnew0 T s : crun (cnew T s : M _). -Proof. by rewrite -(bindskipf (cnew T s)) crunnew // crunskip. Qed. - -Lemma cnewgetret T s : cnew T s >>= @cget _ _ _ _ = cnew T s >> Ret s :> M _. -Proof. under eq_bind do rewrite -cgetret; by rewrite cnewget. Qed. -End cchk. -Arguments cchk {MLU locT M} [T]. - HB.mixin Record isMonadTrace (T : UU0) (M : UU0 -> UU0) of Monad M := { mark : T -> M unit }. diff --git a/impredicative_set/ihierarchy.v b/impredicative_set/ihierarchy.v index a8722b2c..6d1478f0 100644 --- a/impredicative_set/ihierarchy.v +++ b/impredicative_set/ihierarchy.v @@ -283,11 +283,11 @@ End join_laws. End JoinLaws. HB.mixin Record isMonad (F : UU0 -> UU0) of Functor F := { - ret : FId ~> [the functor of F] ; - join : [the functor of F \o F] ~> [the functor of F] ; + ret : FId ~> F ; + join : F \o F ~> F ; bind : forall (A B : UU0), F A -> (A -> F B) -> F B ; - bindE : forall (A B : UU0) (f : A -> F B) (m : F A), - bind A B m f = join B (([the functor of F] # f) m) ; + __bindE : forall (A B : UU0) (f : A -> F B) (m : F A), + bind A B m f = join B ((F # f) m) ; joinretM : JoinLaws.left_unit ret join ; joinMret : JoinLaws.right_unit ret join ; joinA : JoinLaws.associativity join }. @@ -300,6 +300,10 @@ Notation Join := (@join _ _). Arguments bind {s A B} : simpl never. Notation "m >>= f" := (bind m f) : monae_scope. +Lemma bindE (F : monad) (A B : UU0) (f : A -> F B) (m : F A) : + m >>= f = join B ((F # f) m). +Proof. by rewrite __bindE. Qed. + Lemma eq_bind (M : monad) (A B : UU0) (m : M A) (f1 f2 : A -> M B) : f1 =1 f2 -> m >>= f1 = m >>= f2. Proof. by move=> f12; congr bind; apply funext. Qed. @@ -1065,172 +1069,10 @@ HB.structure Definition MonadArray (S : UU0) (I : eqType) := HB.structure Definition MonadPlusArray (S : UU0) (I : eqType) := { M of MonadPlus M & isMonadArray S I M}. -(* A monad for OCaml computations *) -Module Type MLTY. -Parameter ml_type : UU0. -Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Parameter loc : ml_type -> UU0. -Parameter locT : eqType. -Parameter loc_id : forall T, loc T -> locT. -Parameter coq_type : forall M : UU0 -> UU0, ml_type -> UU0. -End MLTY. - -Module MonadTypedStore (MLtypes : MLTY). -Import MLtypes. - -HB.mixin Record isMonadTypedStore (M : UU0 -> UU0) - of Monad M := { - cnew : forall {T}, coq_type M T -> M (loc T) ; - cget : forall {T}, loc T -> M (coq_type M T) ; - cput : forall {T}, loc T -> coq_type M T -> M unit ; - crun : forall {A : UU0}, M A -> option A ; (* execute in empty store *) - cnewget : forall T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A), - cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s) ; - cnewput : forall T (s t : coq_type M T) A (k : loc T -> M A), - cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k ; - cgetput : forall T (r : loc T) (s : coq_type M T), - cget r >> cput r s = cput r s ; - cgetputskip : forall T (r : loc T), cget r >>= cput r = cget r >> skip ; - cgetget : - forall T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A), - cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s ; - cputget : - forall T (r : loc T) (s : coq_type M T) (A : UU0) (k : coq_type M T -> M A), - cput r s >> (cget r >>= k) = cput r s >> k s ; - cputput : forall T (r : loc T) (s s' : coq_type M T), - cput r s >> cput r s' = cput r s' ; - cgetC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) - (k : coq_type M T1 -> coq_type M T2 -> M A), - cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = - cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)) ; - cgetnewD : - forall T T' (r : loc T) (s : coq_type M T') A - (k : loc T' -> coq_type M T -> coq_type M T -> M A), - cget r >>= (fun u => cnew s >>= (fun r' => cget r >>= k r' u)) = - cget r >>= (fun u => cnew s >>= (fun r' => k r' u u)) ; - cgetnewE : forall T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) - (k1 k2 : loc T2 -> M A), - (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> - cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2) ; - cgetputC : forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2), - cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip ; - cputC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) - (s2 : coq_type M T2) (A : UU0), - loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> - cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1 ; - cputgetC : - forall T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) - (A : UU0) (k : coq_type M T2 -> M A), - loc_id r1 != loc_id r2 -> - cput r1 s1 >> cget r2 >>= k = - cget r2 >>= (fun v => cput r1 s1 >> k v) ; - cputnewC : - forall T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A - (k : loc T' -> M A), - cget r >> (cnew s' >>= fun r' => cput r s >> k r') = - cput r s >> (cnew s' >>= k) ; - crunret : forall (A B : UU0) (m : M A) (s : B), - crun m -> crun (m >> Ret s) = Some s ; - crunskip : - crun skip = Some tt ; - crunnew : forall (A : UU0) T (m : M A) (s : coq_type M T), - crun m -> crun (m >> cnew s) ; - }. - -#[short(type=typedStoreMonad)] -HB.structure Definition MonadTypedStore := - { M of isMonadTypedStore M & isMonad M & isFunctor M }. - -Arguments cnew {s}. -Arguments cget {s} [T]. -Arguments cput {s} [T]. -Arguments crun {s} [A]. - -Section cchk. -Variable M : typedStoreMonad. -Definition cchk T (r : loc T) : M unit := cget r >> skip. - -Lemma cnewchk T s (A : UU0) (k : loc T -> M A) : - cnew T s >>= (fun r => cchk r >> k r) = cnew T s >>= k. -Proof. -under eq_bind do rewrite bindA. -rewrite cnewget. -by under eq_bind do rewrite bindskipf. -Qed. - -Lemma cchknewC T1 T2 (r : loc T1) s (A : UU0) (k : loc T2 -> M A) : - cchk r >> (cnew T2 s >>= fun r' => cchk r >> k r') = - cchk r >> (cnew T2 s >>= k). -Proof. -rewrite !(bindA,bindskipf). -under eq_bind do under eq_bind do rewrite bindA. -rewrite cgetnewD. -by under eq_bind do under eq_bind do rewrite bindskipf. -Qed. - -Lemma cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) - (k: coq_type M T2 -> M A) : - cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). -Proof. -under [in RHS]eq_bind do rewrite bindA bindskipf. -by rewrite !(bindA,bindskipf) cgetC. -Qed. - -Lemma cchknewE T1 T2 (r1 : loc T1) s (A : UU0) (k1 k2 : loc T2 -> M A) : - (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> - cchk r1 >> (cnew T2 s >>= k1) = cchk r1 >> (cnew T2 s >>= k2). -Proof. move=> Hk; rewrite !(bindA,bindskipf); exact: cgetnewE. Qed. - -Lemma cchknewget T T' (r : loc T) s (A : UU0) k : - cchk r >> (cnew T' s >>= fun r' => cget r >>= k r') = - cget r >>= (fun u => cnew T' s >>= k ^~ u) :> M A. -Proof. by rewrite bindA bindskipf cgetnewD. Qed. - -Lemma cchknewput T T' (r : loc T) s s' (A : UU0) k : - cchk r >> (cnew T' s' >>= fun r' => cput r s >> k r') = - cput r s >> (cnew T' s' >>= k) :> M A. -Proof. by rewrite bindA bindskipf cputnewC. Qed. - -Lemma cchkget T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : - cchk r >> (cget r >>= k) = cget r >>= k. -Proof. by rewrite bindA bindskipf cgetget. Qed. - -Lemma cgetchk T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : - cget r >>= (fun s => cchk r >> k s) = cget r >>= k. -Proof. under eq_bind do rewrite bindA bindskipf; by rewrite cgetget. Qed. - -Lemma cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : - cchk r1 >> cput r2 s = cput r2 s >> cchk r1. -Proof. by rewrite bindA bindskipf cgetputC bindA. Qed. - -Lemma cchkput T (r : loc T) (s : coq_type M T) : - cchk r >> cput r s = cput r s. -Proof. by rewrite bindA bindskipf cgetput. Qed. - -Lemma cputchk T (r : loc T) (s : coq_type M T) : - cput r s >> cchk r = cput r s. -Proof. by rewrite cputget bindmskip. Qed. - -Lemma cchkC T1 T2 (r1: loc T1) (r2: loc T2) : - cchk r1 >> cchk r2 = cchk r2 >> cchk r1. -Proof. by rewrite !(bindA,bindskipf) cgetC. Qed. - -Lemma cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. -Proof. by rewrite bindA bindskipf cgetget. Qed. - -Lemma cgetret T (r : loc T) : cget r >>= Ret = cget r :> M _. -Proof. by rewrite bindmret. Qed. - -Lemma crunnew0 T s : crun (cnew T s : M _). -Proof. by rewrite -(bindskipf (cnew T s)) crunnew // crunskip. Qed. - -Lemma cnewgetret T s : cnew T s >>= @cget _ _ = cnew T s >> Ret s :> M _. -Proof. under eq_bind do rewrite -cgetret; by rewrite cnewget. Qed. -End cchk. -Arguments cchk {M} [T]. -End MonadTypedStore. +Variant loc (ml_type : Type) (locT : eqType) : ml_type -> Type := + mkloc T : locT -> loc locT T. +Definition loc_id (ml_type : Type) (locT : eqType) {T : ml_type} (l : loc locT T) : locT := + let: mkloc _ n := l in n. HB.mixin Record isMonadTrace (T : UU0) (M : UU0 -> UU0) of Monad M := { mark : T -> M unit }. diff --git a/impredicative_set/imonad_lib.v b/impredicative_set/imonad_lib.v index 6c20f468..f7bbd8fb 100644 --- a/impredicative_set/imonad_lib.v +++ b/impredicative_set/imonad_lib.v @@ -150,6 +150,16 @@ HB.instance Definition _ X := Definition uncurry_F X : functor := [the functor of @uncurry_M X]. End uncurry_functor. +Lemma bind_uncurry (M : monad) (A B C : UU0) (f : A -> M B) (g : A -> B -> M C) x : + (f x >>= fun y => Ret (x, y)) >>= (fun xy => g xy.1 xy.2) = + (f x >>= g x). +Proof. by rewrite bindA; under eq_bind do rewrite bindretf. Qed. + +Lemma bindA_uncurry (M : monad) (A B C : UU0) (m : M A) (f : A -> M B) (g : A -> B -> M C) : + (m >>= fun x => f x >>= fun y => Ret (x, y)) >>= (fun xy => g xy.1 xy.2) = + (m >>= fun x => f x >>= g x). +Proof. by rewrite bindA; by under eq_bind do rewrite bind_uncurry. Qed. + Section exponential_functor. Variable A : UU0. Definition exponential_M (X : UU0) := A -> X. diff --git a/impredicative_set/imonad_model.v b/impredicative_set/imonad_model.v index ddbfba62..24659216 100644 --- a/impredicative_set/imonad_model.v +++ b/impredicative_set/imonad_model.v @@ -1118,542 +1118,6 @@ HB.instance Definition _ := isMonadArray.Build End modelarray. End ModelArray. -Section nth_error. -Context [T : UU0] (def : T) (st : seq T). - -(* Basic lemmas for standard library's nth_error *) -Local Notation nth_error := List.nth_error. - -Lemma nth_error_set_nth n x : - nth_error (set_nth def st n x) n = Some x. -Proof. -elim: n st => [|z IH] [] // {IH}. -elim: z.+1 => [|n <-] //=. -by rewrite set_nth_nil. -Qed. - -Lemma nth_error_rcons_size b : - nth_error (rcons st b) (size st) = Some b. -Proof. by elim: st. Qed. - -Lemma nth_error_rcons_some n a b : - nth_error st n = Some a -> nth_error (rcons st b) n = Some a. -Proof. by elim: n st => [|n IH] []. Qed. - -Lemma nth_error_set_nth_id n a : - nth_error st n = Some a -> set_nth def st n a = st. -Proof. by elim: n st => [|n IH] [] //= b st'; [case=> -> | move/IH ->]. Qed. - -Lemma nth_error_set_nth_other m n a b : - m != n -> - nth_error st m = Some a -> - nth_error (set_nth def st n b) m = Some a. -Proof. elim: m st n => [|m IH] [|c st'] [|n] //=; rewrite eqSS => *; exact: IH. -Qed. - -Lemma nth_error_set_nth_none m n a b : - nth_error st m = None -> - nth_error st n = Some a -> - nth_error (set_nth def st n b) m = None. -Proof. by elim: m st n => [|m IH] [|c st'] [|n] //=; apply IH. Qed. - -Lemma nth_error_size n a : - nth_error st n = Some a -> n < size st. -Proof. by elim: n st => [|n IH] [|c st'] //= /IH. Qed. - -Lemma nth_error_size_set_nth n a b : - nth_error st n = Some a -> size (set_nth def st n b) = size st. -Proof. by rewrite size_set_nth => /nth_error_size /maxn_idPr. Qed. - -Lemma set_nth_rcons a b : - set_nth def (rcons st a) (size st) b = rcons st b. -Proof. by elim: st => //= c st' ->. Qed. - -Lemma nth_error_set_nth_rcons n a b c : - nth_error st n = Some a -> - set_nth def (rcons st c) n b = rcons (set_nth def st n b) c. -Proof. by elim: n st => [|n IH] [|d st'] //= /IH ->. Qed. -End nth_error. - -Require Import JMeq. - -Module Type MLTYweak. -Parameter ml_type : Set. -Parameter ml_type_eq_dec : forall x y : ml_type, {x=y}+{x<>y}. -Variant loc : ml_type -> Set := mkloc T : nat -> loc T. -Parameter coq_type0 : ml_type -> UU0. -Parameter ml_undef : ml_type. -Parameter undef : coq_type0 ml_undef. -End MLTYweak. - -Module ModelTypedStore (MLtypes : MLTYweak). -Import MLtypes. -Module MLtypes'. -Definition ml_type := ml_type. -Definition ml_type_eq_dec := ml_type_eq_dec. -Definition coq_type (M : UU0 -> UU0) := coq_type0. -Definition loc := loc. -Definition locT := [eqType of nat]. -Definition loc_id {T} (l : loc T) := let: mkloc _ n := l in n. -End MLtypes'. -Module MTypedStore := MonadTypedStore (MLtypes'). -Import MLtypes'. -Import MLtypes. -Import MTypedStore. - -Record binding (M : UU0 -> UU0) := - mkbind { bind_type : ml_type; bind_val : coq_type M bind_type }. -Arguments mkbind {M}. - -Section typed_store. -Definition acto (T : UU0) : UU0 := - MS (seq (binding idfun)) [the monad of option_monad] T. -Local Notation M := acto. -(*Local Notation coq_type := (coq_type M).*) -Local Notation coq_type' := (coq_type idfun). - -Definition def : binding idfun := mkbind ml_undef undef. -Local Notation nth_error := List.nth_error. - -Definition cnew T (v : coq_type M T) : M (loc T) := - fun st => let n := size st in - Ret (mkloc T n, rcons st (mkbind T (v : coq_type' T))). - -Definition coerce T1 T2 (v : coq_type M T1) : option (coq_type M T2) := - match ml_type_eq_dec T1 T2 with - | left H => Some (eq_rect _ _ v _ H) - | right _ => None - end. - -Definition cget T (r : loc T) : M (coq_type M T) := - fun st => - if nth_error st (loc_id r) is Some (mkbind T' v) then - if coerce T v is Some u then Ret (u, st) else fail - else fail. - -Definition cput T (r : loc T) (v : coq_type M T) : M unit := - fun st => - let n := loc_id r in - if nth_error st n is Some (mkbind T' _) then - if coerce T' v is Some u then - Ret (tt, set_nth def st n (mkbind T' (u : coq_type' _))) - else fail - else fail. - -Definition crun (A : UU0) (m : M A) : option A := - match m nil with - | inl _ => None - | inr (a, _) => Some a - end. - -(* Make ml_type an eqType *) -Definition ml_type_eqb T1 T2 : bool := ml_type_eq_dec T1 T2. -Lemma ml_type_eqP : Equality.axiom ml_type_eqb. -Proof. -rewrite /ml_type_eqb => T1 T2. -by case: ml_type_eq_dec; constructor. -Qed. -Definition ml_type_eq_mixin := EqMixin ml_type_eqP. -Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. - -(* Prove the laws *) -Let cnewget T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A) : - cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). -Proof. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite /cget nth_error_rcons_size /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq. -Qed. - -Let cnewput T (s t : coq_type M T) A (k : loc T -> M A) : - cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. -Proof. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE bindA. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= [in RHS]bindE /= /cput. -rewrite nth_error_rcons_size. -rewrite /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq bindE /= bindE /= set_nth_rcons. -Qed. - -Let cgetput T (r : loc T) (s : coq_type M T) : - cget r >> cput r s = cput r s. -Proof. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. -case Hr: (nth_error st (loc_id r)) => [[T' u]|] //. -rewrite /coerce. -case: ml_type_eq_dec => HT; last by case: ml_type_eq_dec => // /esym //. -subst T'. -rewrite -eq_rect_eq. -case: ml_type_eq_dec => // HT. -rewrite -eq_rect_eq {HT}. -do! rewrite bindE /=. -rewrite Hr. -case: ml_type_eq_dec => // HT. -by rewrite -eq_rect_eq. -Qed. - -Let cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. -Proof. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE bindA /=. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite /coerce. -case: (ml_type_eq_dec T1 T) => H //. -subst T1. -rewrite bindE /= bindE /= Hr. -case: (ml_type_eq_dec T T) => H //. -by rewrite -eq_rect_eq nth_error_set_nth_id. -Qed. - -Let cgetget T (r : loc T) (A : UU0) (k : coq_type M T -> coq_type M T -> M A) : - cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. -Proof. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite /coerce. -case: (ml_type_eq_dec T1 T) => H //. -subst T1. -rewrite bindE /= bindE /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite Hr. -case: (ml_type_eq_dec T T) => H //. -by rewrite -eq_rect_eq. -Qed. - -Let cputget T (r: loc T) (s: coq_type M T) (A: UU0) (k: coq_type M T -> M A) : - cput r s >> (cget r >>= k) = cput r s >> k s. -Proof. -apply/funext => st /=. -case: r s k => {}T n s k /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. -case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T'. -rewrite !bindretf /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite nth_error_set_nth. -rewrite /coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -by rewrite -eq_rect_eq !bindretf. -Qed. - -Let cputput T (r : loc T) (s s' : coq_type M T) : - cput r s >> cput r s' = cput r s'. -Proof. -apply/funext => st. -case: r s s' => {}T n s s' /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. -case Hst : (nth_error st n) => [[T' v]|] /=; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T'. -rewrite !bindretf /= nth_error_set_nth /coerce. -case: ml_type_eq_dec => // H. -rewrite -eq_rect_eq. -by rewrite set_set_nth eqxx. -Qed. - -Let cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) - (k : coq_type M T1 -> coq_type M T2 -> M A) : - cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = - cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). -Proof. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr1: (nth_error _ _) => [[T1' u]|]; last first. - rewrite bindfailf. - case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. - case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. - rewrite !bindE /=. - by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. -case Hr2: (nth_error _ _) => [[T2' v]|]; last first. - rewrite bindfailf. - case Hc: (coerce T1 u) => [v|]; last by rewrite bindfailf. - rewrite !bindE /=. - by rewrite !bindE /= /bindS MS_mapE /= fmapE Hr2 bindfailf. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec _ _) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE /=. - rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr1 {1}/coerce. - case: (ml_type_eq_dec _ _) => //; by rewrite bindfailf. -subst T1'. -rewrite -eq_rect_eq. -rewrite bindE /=. -rewrite 2!bindE /= /bindS MS_mapE /= fmapE /= bindA /= Hr2 {1}/coerce. -case: (ml_type_eq_dec T2' T2) => HT2; last by rewrite !bindfailf. -subst T2'. -rewrite -eq_rect_eq. -rewrite !bindE /=. -rewrite !bindE /=. -rewrite /bindS MS_mapE /= fmapE /= bindA /= Hr1 /coerce. -case: (ml_type_eq_dec T1 T1) => // H. -by rewrite -eq_rect_eq. -Qed. - -Let cgetnewD T T' (r : loc T) (s : coq_type M T') A - (k : loc T' -> coq_type M T -> coq_type M T -> M A) : - cget r >>= (fun u => cnew s >>= fun r' => cget r >>= k r' u) = - cget r >>= (fun u => cnew s >>= fun r' => k r' u u). -Proof. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cchk /cget. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by rewrite bindfailf. -subst T1. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite (nth_error_rcons_some _ Hr). -rewrite /coerce. -case: ml_type_eq_dec => // H. -by rewrite -eq_rect_eq. -Qed. - -Let cgetnewE T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) - (k1 k2 : loc T2 -> M A) : - (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> - cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2). -Proof. -move=> Hk. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] //=. -case Hc: (coerce T1 v1) => [u1|] //. -rewrite bindE /= bindE /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= [in RHS]bindE /= [in RHS]bindE /= /bindS MS_mapE /=. -rewrite fmapE /= bindA /= 2!bindE /=. -by rewrite Hk // neq_ltn /= (nth_error_size Hr1). -Qed. - -Let cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : - cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip. -Proof. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= 2!bindA /= /cget /cput. -case Hr1 : (nth_error st (loc_id r1)) => [[T1' v1]|] /=; last first. - rewrite bindfailf. - case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; - last by rewrite bindfailf. - case Hc: (coerce T2' s) => [u|]; last by rewrite bindfailf. - rewrite bindE /= /bindS /= bindE /=. - by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). -case Hr2 : (nth_error st (loc_id r2)) => [[T2' v2]|] /=; last first. - rewrite bindfailf. - case Hc: (coerce T1 v1) => [u|]; last by rewrite bindfailf. - by rewrite bindE/= bindE/= Hr2. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec T1' T1) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec T2 T2') => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE/= bindE/=. - case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. - rewrite -Hr nth_error_set_nth /coerce. - move: Hr2; rewrite -Hr Hr1 => -[] HT2. - subst T1'. - by case: ml_type_eq_dec. - rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. - by case: ml_type_eq_dec. -rewrite bindE /= bindE /= Hr2 {1}/coerce. -case: ml_type_eq_dec => HT2 //. -subst T1' T2'. -rewrite -!eq_rect_eq. -rewrite bindE /= bindE /=. -case/boolP: (loc_id r1 == loc_id r2) => [/eqP|] Hr. - rewrite -Hr nth_error_set_nth /coerce. - move: Hr2; rewrite -Hr Hr1 => -[] HT2. - subst T2. - by case: ml_type_eq_dec. -rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. -by case: ml_type_eq_dec. -Qed. - -Let cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) - (s2 : coq_type M T2) (A : UU0) : - loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> - cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. -Proof. -move=> H. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput. -case Hr1: (nth_error _ _) => [[T1' u]|]; last first. - rewrite bindfailf. - case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. - case Hc: (coerce T2' s2) => [u|]; last by rewrite bindfailf. - rewrite !bindE /=. - rewrite !bindE /=. - case/boolP: (loc_id r1 == loc_id r2) => Hr. - by rewrite -(eqP Hr) Hr1 in Hr2. - by rewrite (nth_error_set_nth_none _ _ Hr1 Hr2). -case/boolP: (loc_id r1 == loc_id r2) H => [/eqP|] Hr /= H. - rewrite -Hr Hr1. - case/boolP: (T1 == T2) => /eqP HT; last first. - rewrite /coerce. - case: (ml_type_eq_dec T1 T1') => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec T2 T1') => HT2; last by rewrite bindfailf. - subst T1'. - rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. - by case: ml_type_eq_dec. - subst T1'. - rewrite -eq_rect_eq bindE /= bindE /= nth_error_set_nth. - by case: ml_type_eq_dec => [/esym|] //. - subst T2. - rewrite /coerce. - case: ml_type_eq_dec => HT1 //. - subst T1'. - rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /= !nth_error_set_nth. - case: ml_type_eq_dec => HT1 //. - rewrite -!eq_rect_eq. - case: H => // H. - by move: (JMeq_eq H) => <-. -case Hr2: (nth_error _ _) => [[T2' v]|]; last first. - rewrite bindfailf. - case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. - rewrite !bindE /=. - rewrite !bindE /=. - by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1). -rewrite {1 3}/coerce. -case: (ml_type_eq_dec _ _) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE /=. - rewrite !bindE /=. - rewrite (nth_error_set_nth_other _ _ Hr Hr1) /coerce. - by case: ml_type_eq_dec. -subst T1'. -rewrite -eq_rect_eq /coerce. -rewrite bindE /= bindE /= (nth_error_set_nth_other _ _ _ Hr2); - last by rewrite eq_sym. -case: ml_type_eq_dec => HT2' //. -subst T2'. -rewrite -eq_rect_eq bindE /= bindE /= (nth_error_set_nth_other _ _ Hr Hr1). -case: ml_type_eq_dec => // HT1. -by rewrite -eq_rect_eq set_set_nth (negbTE Hr). -Qed. - -Let cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) - (A : UU0) (k : coq_type M T2 -> M A) : - loc_id r1 != loc_id r2 -> - cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). -Proof. -move=> Hr. -apply/funext => st /=. -rewrite !bindA. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cput /cget. -case Hr1: (nth_error _ _) => [[T1' u]|]; last first. - rewrite bindfailf. - case Hr2: (nth_error _ _) => [[T2' v]|]; last by rewrite bindfailf. - case Hc: (coerce T2 v) => [u|]; last by rewrite bindfailf. - by rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE Hr1 bindfailf. -case Hr2: (nth_error _ _) => [[T2' v]|]; last first. - rewrite bindfailf. - case Hc: (coerce T1' s1) => [v|]; last by rewrite bindfailf. - rewrite !bindE /= !bindE /= /bindS MS_mapE /= fmapE. - by rewrite (nth_error_set_nth_none _ _ Hr2 Hr1) bindfailf. -rewrite {1 3}/coerce. -case: (ml_type_eq_dec _ _) => HT1; last first. - rewrite bindfailf. - case: (ml_type_eq_dec _ _) => HT2; last by rewrite bindfailf. - subst T2'. - rewrite -eq_rect_eq. - rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE Hr1 /coerce. - by case: ml_type_eq_dec. -subst T1'. -rewrite -eq_rect_eq /coerce. -rewrite bindE /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE. -rewrite (nth_error_set_nth_other _ _ _ Hr2); last by rewrite eq_sym. -case: ml_type_eq_dec => HT2' //. -subst T2'. -rewrite -eq_rect_eq bindE /= bindE /= bindE /= bindE /=. -rewrite /bindS /= MS_mapE /= fmapE Hr1. -case: ml_type_eq_dec => // HT1. -by rewrite -eq_rect_eq bindE. -Qed. - -Let cputnewC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A - (k : loc T' -> M A) : - cget r >> (cnew s' >>= fun r' => cput r s >> k r') = - cput r s >> (cnew s' >>= k). -Proof. -apply/funext => st /=. -rewrite bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -rewrite [in RHS]bindE /= /bindS MS_mapE /= fmapE /= bindA /= /cget /cput. -case Hr: (nth_error _ _) => [[T1 u]|]; last by rewrite bindfailf. -rewrite {1 3}/coerce. -case: ml_type_eq_dec => H /=; last by case ml_type_eq_dec => H' //; elim H. -subst T1. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite (nth_error_rcons_some _ Hr). -rewrite /coerce. -case: ml_type_eq_dec => // H. -rewrite -eq_rect_eq. -rewrite bindE /= /bindS /= bindE /= bindE /= /bindS /= MS_mapE /= fmapE /=. -rewrite bindE /= /bindS /=. -rewrite (nth_error_size_set_nth _ _ Hr). -by rewrite (nth_error_set_nth_rcons _ _ _ Hr). -Qed. - -Let crunret (A B : UU0) (m : M A) (s : B) : - crun m -> crun (m >> Ret s) = Some s. -Proof. -rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -by case Hm: (m [::]). -Qed. - -Let crunskip : crun skip = Some tt. -Proof. by []. Qed. - -Let crunnew (A : UU0) T (m : M A) (s : coq_type M T) : - crun m -> crun (m >> cnew s). -Proof. -rewrite /crun /= bindE /= /bindS MS_mapE /= fmapE /= bindA /=. -by case Hm: (m [::]). -Qed. - -Canonical Structure isMonadTypedStoreModel := - isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip - cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC - cputgetC cputnewC crunret crunskip crunnew. - -(* The elpi tactic/command HB.instance failed without giving a specific - error message. -HB.instance Definition _ := - isMonadTypedStore.Build M cnewget cnewput cgetput cgetputskip - cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC - cputgetC cputnewC crunret crunskip crunnew. -*) -End typed_store. -End ModelTypedStore. - (* TODO? (* result of a discussion with Maxime and Enrico on 2019-09-12 *) (* Equality between monads from the hierarchy and their counterparts built *) diff --git a/monad_lib.v b/monad_lib.v index fb22097b..389ef16c 100644 --- a/monad_lib.v +++ b/monad_lib.v @@ -153,13 +153,13 @@ HB.instance Definition _ X := Definition uncurry_F X : functor := [the functor of @uncurry_M X]. End uncurry_functor. -Lemma bind_uncurry (M : monad) A B C (f : A -> M B) (g : A -> B -> M C) x : - (f x >>= fun y => Ret (x,y)) >>= (fun xy => g xy.1 xy.2)(*uncurry g*) = +Lemma bind_uncurry (M : monad) (A B C : UU0) (f : A -> M B) (g : A -> B -> M C) x : + (f x >>= fun y => Ret (x, y)) >>= (fun xy => g xy.1 xy.2) = (f x >>= g x). Proof. by rewrite bindA; under eq_bind do rewrite bindretf. Qed. -Lemma bindA_uncurry (M : monad) A B C (m : M A) (f : A -> M B) (g : A -> B -> M C) : - (m >>= fun x => f x >>= fun y => Ret (x,y)) >>= (fun xy => g xy.1 xy.2) = +Lemma bindA_uncurry (M : monad) (A B C : UU0) (m : M A) (f : A -> M B) (g : A -> B -> M C) : + (m >>= fun x => f x >>= fun y => Ret (x, y)) >>= (fun xy => g xy.1 xy.2) = (m >>= fun x => f x >>= g x). Proof. by rewrite bindA; by under eq_bind do rewrite bind_uncurry. Qed. @@ -596,7 +596,7 @@ Proof. by []. Qed. Lemma naturality_mpair (M : monad) (A B : UU0) (f : A -> B) (g : A -> M A): (M # f^`2) \o (mpair \o g^`2) = mpair \o ((M # f) \o g)^`2. Proof. -rewrite boolp.funeqE => -[a0 a1]. +apply boolp.funext => -[a0 a1]. rewrite compE fmap_bind. rewrite compE mpairE compE bind_fmap; bind_ext => a2. rewrite fcompE fmap_bind 2!compE bind_fmap; bind_ext => a3. diff --git a/typed_store_lib.v b/typed_store_lib.v new file mode 100644 index 00000000..633e6411 --- /dev/null +++ b/typed_store_lib.v @@ -0,0 +1,99 @@ +(* monae: Monadic equational reasoning in Coq *) +(* Copyright (C) 2020 monae authors, license: LGPL-2.1-or-later *) +Ltac typeof X := type of X. + +Require Import ssrmatching Reals JMeq. +From mathcomp Require Import all_ssreflect. +From mathcomp Require boolp. +From infotheo Require Import Reals_ext. +Require Import monae_lib hierarchy. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Local Open Scope monae_scope. + +Section cchk. +Variables (MLU : ML_universe) (locT : eqType) (M : typedStoreMonad MLU locT). +Notation loc := (@loc MLU locT). +Definition cchk T (r : loc T) : M unit := cget r >> skip. + +Lemma cnewchk T s (A : UU0) (k : loc T -> M A) : + cnew T s >>= (fun r => cchk r >> k r) = cnew T s >>= k. +Proof. +under eq_bind do rewrite bindA. +rewrite cnewget. +by under eq_bind do rewrite bindskipf. +Qed. + +Lemma cchknewC T1 T2 (r : loc T1) s (A : UU0) (k : loc T2 -> M A) : + cchk r >> (cnew T2 s >>= fun r' => cchk r >> k r') = + cchk r >> (cnew T2 s >>= k). +Proof. +rewrite !(bindA,bindskipf). +under eq_bind do under eq_bind do rewrite bindA. +rewrite cgetnewD. +by under eq_bind do under eq_bind do rewrite bindskipf. +Qed. + +Lemma cchkgetC T1 T2 (r1: loc T1) (r2: loc T2) (A: UU0) + (k: coq_type M T2 -> M A) : + cchk r1 >> (cget r2 >>= k) = cget r2 >>= (fun s => cchk r1 >> k s). +Proof. +under [in RHS]eq_bind do rewrite bindA bindskipf. +by rewrite !(bindA,bindskipf) cgetC. +Qed. + +Lemma cchknewE T1 T2 (r1 : loc T1) s (A : UU0) (k1 k2 : loc T2 -> M A) : + (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> + cchk r1 >> (cnew T2 s >>= k1) = cchk r1 >> (cnew T2 s >>= k2). +Proof. move=> Hk; rewrite !(bindA,bindskipf); exact: cgetnewE. Qed. + +Lemma cchknewget T T' (r : loc T) s (A : UU0) k : + cchk r >> (cnew T' s >>= fun r' => cget r >>= k r') = + cget r >>= (fun u => cnew T' s >>= k ^~ u) :> M A. +Proof. by rewrite bindA bindskipf cgetnewD. Qed. + +Lemma cchknewput T T' (r : loc T) s s' (A : UU0) k : + cchk r >> (cnew T' s' >>= fun r' => cput r s >> k r') = + cput r s >> (cnew T' s' >>= k) :> M A. +Proof. by rewrite bindA bindskipf cputnewC. Qed. + +Lemma cchkget T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : + cchk r >> (cget r >>= k) = cget r >>= k. +Proof. by rewrite bindA bindskipf cgetget. Qed. + +Lemma cgetchk T (r : loc T) (A: UU0) (k : coq_type M T -> M A) : + cget r >>= (fun s => cchk r >> k s) = cget r >>= k. +Proof. under eq_bind do rewrite bindA bindskipf; by rewrite cgetget. Qed. + +Lemma cchkputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : + cchk r1 >> cput r2 s = cput r2 s >> cchk r1. +Proof. by rewrite bindA bindskipf cgetputC bindA. Qed. + +Lemma cchkput T (r : loc T) (s : coq_type M T) : + cchk r >> cput r s = cput r s. +Proof. by rewrite bindA bindskipf cgetput. Qed. + +Lemma cputchk T (r : loc T) (s : coq_type M T) : + cput r s >> cchk r = cput r s. +Proof. by rewrite cputget bindmskip. Qed. + +Lemma cchkC T1 T2 (r1: loc T1) (r2: loc T2) : + cchk r1 >> cchk r2 = cchk r2 >> cchk r1. +Proof. by rewrite !(bindA,bindskipf) cgetC. Qed. + +Lemma cchkdup T (r : loc T) : cchk r >> cchk r = cchk r. +Proof. by rewrite bindA bindskipf cgetget. Qed. + +Lemma cgetret T (r : loc T) : cget r >>= Ret = cget r :> M _. +Proof. by rewrite bindmret. Qed. + +Lemma crunnew0 T s : crun (cnew T s : M _). +Proof. by rewrite -(bindskipf (cnew T s)) crunnew // crunskip. Qed. + +Lemma cnewgetret T s : cnew T s >>= @cget _ _ _ _ = cnew T s >> Ret s :> M _. +Proof. under eq_bind do rewrite -cgetret; by rewrite cnewget. Qed. +End cchk. +Arguments cchk {MLU locT M} [T]. From 848dbeb2ef1cda54f13188b30b0b4032ab7617c9 Mon Sep 17 00:00:00 2001 From: Jacques Garrigue Date: Tue, 25 Jul 2023 11:11:31 +0900 Subject: [PATCH 78/82] shared ml_type_eq_mixin --- example_typed_store.v | 9 +++------ 1 file changed, 3 insertions(+), 6 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 956c9591..1726a51f 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -54,6 +54,9 @@ Inductive rlist (a : Type) (a_1 : ml_type) := | Nil | Cons (_ : a) (_ : loc (ml_rlist a_1)). +Definition ml_type_eq_mixin := comparableMixin MLTypes.ml_type_eq_dec. +Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. + End MLTypes. Module MLTypesNat. @@ -74,9 +77,6 @@ Fixpoint coq_type_nat (T : ml_type) : Type := end. End with_monad. -Definition ml_type_eq_mixin := comparableMixin MLTypes.ml_type_eq_dec. -Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. - (*Canonical MLU := @Build_ML_universe _ coq_type ml_unit val_nonempty.*) Definition coq_type0 := @coq_type_nat idfun. @@ -452,9 +452,6 @@ Fixpoint coq_type63 (T : ml_type) : Type := end. End with_monad. -Definition ml_type63_eq_mixin := comparableMixin MLTypes.ml_type_eq_dec. -Canonical ml_type63_eqType := Eval hnf in EqType _ ml_type63_eq_mixin. - Canonical ml_type63 := @Build_ML_universe _ coq_type63 ml_unit val_nonempty. Section fact_for_int63. From 393ab5e9d599a40ae08b70cd734a63e7559f1b6b Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 25 Jul 2023 11:11:47 +0900 Subject: [PATCH 79/82] rm do notation --- example_typed_store.v | 11 +++-------- 1 file changed, 3 insertions(+), 8 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 1726a51f..7d7b4e4c 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -85,8 +85,7 @@ Canonical MLU := @monad_model.MLU _ coq_type0 ml_unit (val_nonempty idfun). Section cyclic. Variables (M : typedStoreMonad MLU MLTypes.locT). Notation coq_type := (@hierarchy.coq_type MLU M). -Notation "'do' x <- m ; e" := (m >>= (fun x => e)) - (at level 60, x name, m at level 200, e at level 60). +Local Open Scope do_notation. Definition cycle (T : ml_type) (a b : coq_type T) : M (coq_type (ml_rlist T)) := @@ -186,9 +185,7 @@ Arguments forloop {M}. Section fact_for. Variable M : typedStoreMonad MLU MLTypes.locT. Notation coq_type := (@coq_type MLU M). - -Notation "'do' x <- m ; e" := (m >>= (fun x => e)) - (at level 60, x name, m at level 200, e at level 60). +Local Open Scope do_notation. Definition fact_for (n : coq_type ml_int) : M (coq_type ml_int) := do v <- cnew ml_int 1; @@ -457,9 +454,7 @@ Canonical ml_type63 := @Build_ML_universe _ coq_type63 ml_unit val_nonempty. Section fact_for_int63. Variable M : typedStoreMonad ml_type63 MLTypes.locT. Notation coq_type := (@MLtypes63.coq_type63 M). - -Notation "'do' x <- m ; e" := (m >>= (fun x => e)) - (at level 60, x name, m at level 200, e at level 60). +Local Open Scope do_notation. Section forloop63. Definition forloop63 (n_1 n_2 : int) (b : int -> M unit) : M unit := From 566280291656ac8be3a27eb10196f805db01ddab Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 25 Jul 2023 14:50:26 +0900 Subject: [PATCH 80/82] nitpick --- example_typed_store.v | 72 ++++--------- impredicative_set/imonad_lib.v | 38 +++++++ impredicative_set/imonae_lib.v | 70 +++++++++++- monad_lib.v | 38 +++++++ monad_model.v | 192 +++++++++------------------------ monad_transformer.v | 4 + monae_lib.v | 97 ++++++++++++++++- typed_store_lib.v | 6 +- typed_store_model.v | 31 +++--- 9 files changed, 331 insertions(+), 217 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 7d7b4e4c..999c33a5 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -1,5 +1,5 @@ (* monae: Monadic equational reasoning in Coq *) -(* Copyright (C) 2020 monae authors, license: LGPL-2.1-or-later *) +(* Copyright (C) 2023 monae authors, license: LGPL-2.1-or-later *) Require Import ZArith. From mathcomp Require Import all_ssreflect. From mathcomp Require boolp. @@ -10,20 +10,26 @@ Require Import monae_lib hierarchy monad_lib typed_store_lib. (******************************************************************************) (* Typed store examples *) (* *) -(* Inductive ml_type *) +(* Inductive ml_type == generated by coqgen *) (* *) (* Module MLTypesNat *) +(* coq_type_nat == adapted from code generated by coqgen *) +(* coq_type_nat0 == coq_type_nat with identity monad *) (* Definition cycle *) (* Fixpoint fact_ref *) (* Definition fact_for *) (* Fixpoint fibo_ref *) (* *) (* Module MLtypes63 *) +(* Fixpoint coq_type63 == generated type translation function *) (* Definition fact_for63 *) (******************************************************************************) Local Open Scope monae_scope. +(******************************************************************************) +(* generated by coqgen *) +(******************************************************************************) Module MLTypes. Inductive ml_type : Set := | ml_int @@ -58,6 +64,7 @@ Definition ml_type_eq_mixin := comparableMixin MLTypes.ml_type_eq_dec. Canonical ml_type_eqType := Eval hnf in EqType _ ml_type_eq_mixin. End MLTypes. +(******************************************************************************) Module MLTypesNat. Import MLTypes. @@ -65,7 +72,6 @@ Import MLTypes. Section with_monad. Context [M : Type -> Type]. -(* Generated type translation function *) Fixpoint coq_type_nat (T : ml_type) : Type := match T with | ml_int => nat @@ -77,10 +83,14 @@ Fixpoint coq_type_nat (T : ml_type) : Type := end. End with_monad. -(*Canonical MLU := @Build_ML_universe _ coq_type ml_unit val_nonempty.*) +(* use coq_type_nat (typed_store_model.v) *) +Canonical MLU := @Build_ML_universe _ coq_type_nat ml_unit val_nonempty. -Definition coq_type0 := @coq_type_nat idfun. -Canonical MLU := @monad_model.MLU _ coq_type0 ml_unit (val_nonempty idfun). +(* alternative: use coq_type_nat0 (monad_model.v) *) +(* +Definition coq_type_nat0 := @coq_type_nat idfun. +Canonical MLU := @monad_model.MLU _ coq_type_nat0 ml_unit (val_nonempty idfun). +*) Section cyclic. Variables (M : typedStoreMonad MLU MLTypes.locT). @@ -143,45 +153,6 @@ by rewrite cnewput IH // (mulnC m.+1) -mulnA. Qed. End factorial. -Section forloop. -Variable M : monad. - -Definition forloop (n_1 n_2 : nat) (b : nat -> M unit) : M unit := - if n_2 < n_1 then Ret tt else - iteri (n_2.+1 - n_1) - (fun i (m : M unit) => m >> b (n_1 + i)) - skip. - -Lemma iteriSr T n (f : nat -> T -> T) x : - iteri n.+1 f x = iteri n (f \o succn) (f 0 x). -Proof. by elim: n x => // n IH x /=; rewrite -IH. Qed. - -Lemma iteriD T n m (f : nat -> T -> T) x : - iteri (n + m) f x = iteri m (f \o addn n) (iteri n f x). -Proof. by elim: n x f => // n IH x f; rewrite addSn iteriSr IH iteriSr. Qed. - -Lemma iteri_bind n (f : nat -> M unit) (m1 m2 : M unit) : - iteri n (fun i (m : M unit) => m >> f i) (m1 >> m2) = - m1 >> iteri n (fun i (m : M unit) => m >> f i) m2. -Proof. by elim: n m2 => // n IH m2; rewrite iteriS IH !bindA. Qed. - -Lemma forloopS m n (f : nat -> M unit) : - m <= n -> forloop m n f = f m >> forloop m.+1 n f. -Proof. -rewrite /forloop => mn. -rewrite ltnNge mn /= subSS subSn // iteriSr bindskipf. -rewrite -[f _]bindmskip iteri_bind addn0 ltnS -subn_eq0. -case: (n-m) => //= k. -rewrite addSnnS; apply eq_bind => _; congr bind. -apply eq_iteri => i x; by rewrite addSnnS. -Qed. - -Lemma forloop0 m n (f : nat -> M unit) : - m > n -> forloop m n f = skip. -Proof. by rewrite /forloop => ->. Qed. -End forloop. -Arguments forloop {M}. - Section fact_for. Variable M : typedStoreMonad MLU MLTypes.locT. Notation coq_type := (@coq_type MLU M). @@ -434,10 +405,11 @@ End Int63. Module MLtypes63. Import MLTypes. +(******************************************************************************) +(* generated by coqgen *) +(******************************************************************************) Section with_monad. Context [M : Type -> Type]. - -(* Generated type translation function *) Fixpoint coq_type63 (T : ml_type) : Type := match T with | ml_int => int @@ -448,6 +420,7 @@ Fixpoint coq_type63 (T : ml_type) : Type := | ml_rlist T1 => rlist (coq_type63 T1) T1 end. End with_monad. +(******************************************************************************) Canonical ml_type63 := @Build_ML_universe _ coq_type63 ml_unit val_nonempty. @@ -463,11 +436,6 @@ Definition forloop63 (n_1 n_2 : int) (b : int -> M unit) : M unit := (fun (m : M int) => do i <- m; do _ <- b i; Ret (Uint63.succ i)) (Ret n_1) >> Ret tt. -Lemma iter_bind T n (f : T -> M T) (m1 : M unit) m2 : - iter n (fun (m : M T) => m >>= f) (m1 >> m2) = - m1 >> iter n (fun (m : M T) => m >>= f) m2. -Proof. by elim: n m2 => // n IH m2; rewrite iterS IH !bindA. Qed. - Lemma forloop63S m n (f : int -> M unit) : ltsb m n -> forloop63 m n f = f m >> forloop63 (Uint63.succ m) n f. Proof. diff --git a/impredicative_set/imonad_lib.v b/impredicative_set/imonad_lib.v index f7bbd8fb..c77c41c4 100644 --- a/impredicative_set/imonad_lib.v +++ b/impredicative_set/imonad_lib.v @@ -34,6 +34,7 @@ Require Import ihierarchy. (* commute m n f := (m >>= n >>= f) = (n >>= m >>= f) *) (* (ref: definition 4.2, mu2019tr3) *) (* rep n mx == mx >> mx >> ... >> mx, n times *) +(* forloop n1 n2 (b : nat -> M unit) : M unit := for-loop *) (* *) (******************************************************************************) @@ -706,3 +707,40 @@ by rewrite -subn1 addnBA ?expn_gt0 // addnn -muln2 -expnSr subn1. Qed. End MonadCount. + +Lemma iteri_bind {M : monad} n (f : nat -> M unit) (m1 m2 : M unit) : + iteri n (fun i (m : M unit) => m >> f i) (m1 >> m2) = + m1 >> iteri n (fun i (m : M unit) => m >> f i) m2. +Proof. by elim: n m2 => // n IH m2; rewrite iteriS IH !bindA. Qed. + +Lemma iter_bind {M : monad} (T : UU0) n (f : T -> M T) (m1 : M unit) m2 : + iter n (fun (m : M T) => m >>= f) (m1 >> m2) = + m1 >> iter n (fun (m : M T) => m >>= f) m2. +Proof. by elim: n m2 => // n IH m2; rewrite iterS IH !bindA. Qed. + +Section forloop. +Variable M : monad. + +Definition forloop (n_1 n_2 : nat) (b : nat -> M unit) : M unit := + if n_2 < n_1 then Ret tt else + iteri (n_2.+1 - n_1) + (fun i (m : M unit) => m >> b (n_1 + i)) + skip. + +Lemma forloopS m n (f : nat -> M unit) : + m <= n -> forloop m n f = f m >> forloop m.+1 n f. +Proof. +rewrite /forloop => mn. +rewrite ltnNge mn /= subSS subSn // iteriSr bindskipf. +rewrite -[f _]bindmskip iteri_bind addn0 ltnS -subn_eq0. +case: (n-m) => //= k. +rewrite addSnnS; apply eq_bind => _; congr bind. +apply eq_iteri => i x; by rewrite addSnnS. +Qed. + +Lemma forloop0 m n (f : nat -> M unit) : + m > n -> forloop m n f = skip. +Proof. by rewrite /forloop => ->. Qed. + +End forloop. +Arguments forloop {M}. diff --git a/impredicative_set/imonae_lib.v b/impredicative_set/imonae_lib.v index 84c49abb..97bb1c39 100644 --- a/impredicative_set/imonae_lib.v +++ b/impredicative_set/imonae_lib.v @@ -5,7 +5,8 @@ Require ProofIrrelevance FunctionalExtensionality. Definition proof_irr := @ProofIrrelevance.proof_irrelevance. -Definition eq_rect_eq := @ProofIrrelevance.ProofIrrelevanceTheory.Eq_rect_eq.eq_rect_eq. +Definition eq_rect_eq := + @ProofIrrelevance.ProofIrrelevanceTheory.Eq_rect_eq.eq_rect_eq. Definition funext := @FunctionalExtensionality.functional_extensionality. @@ -138,3 +139,70 @@ Lemma compidf A B (f : A -> B) : id \o f = f. Proof. by []. Qed. Lemma compE A B C (g : B -> C) (f : A -> B) a : (g \o f) a = g (f a). Proof. by []. Qed. + +Lemma if_pair A B b (x : A) y (u : A) (v : B) : + (if b then (x, y) else (u, v)) = (if b then x else u, if b then y else v). +Proof. by case: ifPn. Qed. + +Lemma iteriSr T n (f : nat -> T -> T) x : + iteri n.+1 f x = iteri n (f \o succn) (f 0 x). +Proof. by elim: n x => // n IH x /=; rewrite -IH. Qed. + +Lemma iteriD T n m (f : nat -> T -> T) x : + iteri (n + m) f x = iteri m (f \o addn n) (iteri n f x). +Proof. by elim: n x f => // n IH x f; rewrite addSn iteriSr IH iteriSr. Qed. + +Section nth_error. +Context [T : Type] (def : T) (st : seq T). + +(* Basic lemmas for standard library's nth_error *) +Local Notation nth_error := List.nth_error. + +Lemma nth_error_set_nth n x : nth_error (set_nth def st n x) n = Some x. +Proof. +elim: n st => [|z IH] [] // {IH}. +elim: z.+1 => [|n <-] //=. +by rewrite set_nth_nil. +Qed. + +Lemma nth_error_rcons_size b : nth_error (rcons st b) (size st) = Some b. +Proof. by elim: st. Qed. + +Lemma nth_error_rcons_some n a b : + nth_error st n = Some a -> nth_error (rcons st b) n = Some a. +Proof. by elim: n st => [|n IH] []. Qed. + +Lemma nth_error_set_nth_id n a : + nth_error st n = Some a -> set_nth def st n a = st. +Proof. by elim: n st => [|n IH] [] //= b st'; [case=> -> | move/IH ->]. Qed. + +Lemma nth_error_set_nth_other m n a b : + m != n -> + nth_error st m = Some a -> + nth_error (set_nth def st n b) m = Some a. +Proof. +elim: m st n => [|m IH] [|c st'] [|n] //=; rewrite eqSS => *; exact: IH. +Qed. + +Lemma nth_error_set_nth_none m n a b : + nth_error st m = None -> + nth_error st n = Some a -> + nth_error (set_nth def st n b) m = None. +Proof. by elim: m st n => [|m IH] [|c st'] [|n] //=; apply IH. Qed. + +Lemma nth_error_size n a : nth_error st n = Some a -> n < size st. +Proof. by elim: n st => [|n IH] [|c st'] //= /IH. Qed. + +Lemma nth_error_size_set_nth n a b : + nth_error st n = Some a -> size (set_nth def st n b) = size st. +Proof. by rewrite size_set_nth => /nth_error_size /maxn_idPr. Qed. + +Lemma set_nth_rcons a b : set_nth def (rcons st a) (size st) b = rcons st b. +Proof. by elim: st => //= c st' ->. Qed. + +Lemma nth_error_set_nth_rcons n a b c : + nth_error st n = Some a -> + set_nth def (rcons st c) n b = rcons (set_nth def st n b) c. +Proof. by elim: n st => [|n IH] [|d st'] //= /IH ->. Qed. +End nth_error. +Arguments nth_error_size {T st n a}. diff --git a/monad_lib.v b/monad_lib.v index 389ef16c..1968c035 100644 --- a/monad_lib.v +++ b/monad_lib.v @@ -37,6 +37,7 @@ Require Import hierarchy. (* preserves f g := the monadic function f : A -> M A preserves *) (* the value of the function g : A -> B *) (* rep n mx == mx >> mx >> ... >> mx, n times *) +(* forloop n1 n2 (b : nat -> M unit) : M unit := for-loop *) (* *) (******************************************************************************) @@ -709,3 +710,40 @@ by rewrite -subn1 addnBA ?expn_gt0 // addnn -muln2 -expnSr subn1. Qed. End MonadCount. + +Lemma iteri_bind {M : monad} n (f : nat -> M unit) (m1 m2 : M unit) : + iteri n (fun i (m : M unit) => m >> f i) (m1 >> m2) = + m1 >> iteri n (fun i (m : M unit) => m >> f i) m2. +Proof. by elim: n m2 => // n IH m2; rewrite iteriS IH !bindA. Qed. + +Lemma iter_bind {M : monad} (T : UU0) n (f : T -> M T) (m1 : M unit) m2 : + iter n (fun (m : M T) => m >>= f) (m1 >> m2) = + m1 >> iter n (fun (m : M T) => m >>= f) m2. +Proof. by elim: n m2 => // n IH m2; rewrite iterS IH !bindA. Qed. + +Section forloop. +Variable M : monad. + +Definition forloop (n_1 n_2 : nat) (b : nat -> M unit) : M unit := + if n_2 < n_1 then Ret tt else + iteri (n_2.+1 - n_1) + (fun i (m : M unit) => m >> b (n_1 + i)) + skip. + +Lemma forloopS m n (f : nat -> M unit) : + m <= n -> forloop m n f = f m >> forloop m.+1 n f. +Proof. +rewrite /forloop => mn. +rewrite ltnNge mn /= subSS subSn // iteriSr bindskipf. +rewrite -[f _]bindmskip iteri_bind addn0 ltnS -subn_eq0. +case: (n-m) => //= k. +rewrite addSnnS; apply eq_bind => _; congr bind. +apply eq_iteri => i x; by rewrite addSnnS. +Qed. + +Lemma forloop0 m n (f : nat -> M unit) : + m > n -> forloop m n f = skip. +Proof. by rewrite /forloop => ->. Qed. + +End forloop. +Arguments forloop {M}. diff --git a/monad_model.v b/monad_model.v index 23774811..d1b97c39 100644 --- a/monad_model.v +++ b/monad_model.v @@ -1,3 +1,4 @@ +Require Import JMeq. From mathcomp Require Import all_ssreflect. From mathcomp Require Import finmap. From mathcomp Require boolp. @@ -62,6 +63,8 @@ Require Import monad_transformer. (* Module ModelMonadStateRun == stateRunMonad for MS *) (* Module ModelMonadExceptStateRun == exceptStateRunMonad *) (* *) +(* Module ModelTypedStore == model for typed store *) +(* *) (* references: *) (* - Wadler, P. Monads and composable continuations. LISP and Symbolic *) (* Computation 7, 39–55 (1994) *) @@ -1658,127 +1661,37 @@ End modelplusarray. End ModelPlusArray. HB.export ModelPlusArray. -Section nth_error. -Context [T : Type] (def : T) (st : seq T). - -(* Basic lemmas for standard library's nth_error *) -Local Notation nth_error := List.nth_error. - -Lemma nth_error_set_nth n x : nth_error (set_nth def st n x) n = Some x. -Proof. -elim: n st => [|z IH] [] // {IH}. -elim: z.+1 => [|n <-] //=. -by rewrite set_nth_nil. -Qed. - -Lemma nth_error_rcons_size b : nth_error (rcons st b) (size st) = Some b. -Proof. by elim: st. Qed. - -Lemma nth_error_rcons_some n a b : - nth_error st n = Some a -> nth_error (rcons st b) n = Some a. -Proof. by elim: n st => [|n IH] []. Qed. - -Lemma nth_error_set_nth_id n a : - nth_error st n = Some a -> set_nth def st n a = st. -Proof. by elim: n st => [|n IH] [] //= b st'; [case=> -> | move/IH ->]. Qed. +Definition locT_nat := [eqType of nat]. -Lemma nth_error_set_nth_other m n a b : - m != n -> - nth_error st m = Some a -> - nth_error (set_nth def st n b) m = Some a. -Proof. -elim: m st n => [|m IH] [|c st'] [|n] //=; rewrite eqSS => *; exact: IH. -Qed. - -Lemma nth_error_set_nth_none m n a b : - nth_error st m = None -> - nth_error st n = Some a -> - nth_error (set_nth def st n b) m = None. -Proof. by elim: m st n => [|m IH] [|c st'] [|n] //=; apply IH. Qed. - -Lemma nth_error_size n a : nth_error st n = Some a -> n < size st. -Proof. by elim: n st => [|n IH] [|c st'] //= /IH. Qed. - -Lemma nth_error_size_set_nth n a b : - nth_error st n = Some a -> size (set_nth def st n b) = size st. -Proof. by rewrite size_set_nth => /nth_error_size /maxn_idPr. Qed. - -Lemma set_nth_rcons a b : set_nth def (rcons st a) (size st) b = rcons st b. -Proof. by elim: st => //= c st' ->. Qed. - -Lemma nth_error_set_nth_rcons n a b c : - nth_error st n = Some a -> - set_nth def (rcons st c) n b = rcons (set_nth def st n b) c. -Proof. by elim: n st => [|n IH] [|d st'] //= /IH ->. Qed. -End nth_error. -Arguments nth_error_size {T st n a}. - -Require Import JMeq. - -(*Structure ML_universe0 := { - ml_type0 :> eqType ; - coq_type0 : ml_type0 -> Type ; - ml_nonempty0 : ml_type0 ; - val_nonempty0 : coq_type0 ml_nonempty0 }.*) - -Section coerce. -Variables (X : eqType) (f : X -> UU0). - -Definition coerce (T1 T2 : X) (v : f T1) : option (f T2) := - if @eqP _ T1 T2 is ReflectT H then Some (eq_rect _ _ v _ H) else None. - -Lemma coerce_sym (T T' : X) (s : f T) (s' : f T') : coerce T' s -> coerce T s'. -Proof. -by rewrite /coerce; case: eqP => //= h; case: eqP => //; rewrite h; auto. -Qed. - -Lemma coerce_Some (T : X) (s : f T) : coerce T s = Some s. -Proof. -by rewrite /coerce; case: eqP => /= [?|]; [rewrite -eq_rect_eq|auto]. -Qed. - -Lemma coerce_eq (T T' : X) (s : f T) : coerce T' s -> T = T'. -Proof. by rewrite /coerce; case: eqP. Qed. - -Lemma coerce_None (T T' : X) (s : f T) : T != T' -> coerce T' s = None. -Proof. by rewrite /coerce; case: eqP. Qed. - -End coerce. +Module ModelTypedStore. Section ModelTypedStore. -(*Variable MLU0 : ML_universe0.*) -Variables - (ml_type0 : eqType) - (coq_type0 : ml_type0 -> Type) - (ml_nonempty0 : ml_type0) - (val_nonempty0 : coq_type0 ml_nonempty0). +Variables (ml_type0 : eqType) (coq_type0 : ml_type0 -> Type) + (ml_nonempty0 : ml_type0) (val_nonempty0 : coq_type0 ml_nonempty0). Definition MLU : ML_universe := @Build_ML_universe (ml_type0) (fun M => @coq_type0) (ml_nonempty0) (fun M => val_nonempty0). -Definition locT := [eqType of nat]. - Record binding (M : Type -> Type) := mkbind { bind_type : MLU; bind_val : coq_type M bind_type }. Arguments mkbind {M}. -Section typed_store. Definition acto : UU0 -> UU0 := MS (seq (binding idfun)) [the monad of option_monad]. Local Notation M := acto. + Local Notation coq_type' := (@coq_type MLU idfun). Definition def : binding idfun := mkbind (ml_nonempty MLU) (val_nonempty MLU idfun). Local Notation nth_error := List.nth_error. -Notation loc := (@loc MLU locT). +Notation loc := (@loc MLU locT_nat). Definition cnew T (v : coq_type M T) : M (loc T) := fun st => let n := size st in Ret (mkloc T n, rcons st (mkbind T (v : coq_type' T))). - Definition cget T (r : loc T) : M (coq_type M T) := fun st => if nth_error st (loc_id r) is Some (mkbind T' v) then @@ -1802,15 +1715,31 @@ Definition crun (A : UU0) (m : M A) : option A := Definition Env := seq (binding idfun). +Definition fresh_loc (T : MLU) (e : Env) := mkloc T (size e). + +Section mkbind. Local Notation mkbind := (@mkbind idfun). Definition extend_env T (v : coq_type M T) (e : Env) := rcons e (mkbind T v). -Definition fresh_loc (T : MLU) (e : Env) := mkloc T (size e). -Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M B) s : - (m >>= f) s = m s >>= uncurry f. -Proof. by []. Qed. +Variant nth_error_spec (T : MLU) (e : Env) (r : loc T) : Type := + | NthError : forall s : coq_type M T, + nth_error e (loc_id r) = Some (mkbind T s) -> nth_error_spec e r + | NthError_nocoerce : forall T' (s' : coq_type M T'), + nth_error e (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> + nth_error_spec e r + | NthError_None : nth_error e (loc_id r) = None -> nth_error_spec e r. + +Lemma ntherrorP (T : MLU) (e : Env) (r : loc T) : nth_error_spec e r. +Proof. +case H : (nth_error e (loc_id r)) => [[T' s']|]. + have [Ts'|Ts'] := boolp.pselect (coerce T s'). + have ? := coerce_eq Ts'; subst T'. + exact: NthError H. + exact: NthError_nocoerce H Ts'. +exact: NthError_None. +Qed. Lemma bind_cnew T (s : coq_type M T) A (k : loc T -> M A) e : (cnew s >>= k) e = k (fresh_loc T e) (extend_env s e). @@ -1822,22 +1751,14 @@ Lemma Some_cget T (r : loc T) (s : coq_type M T) e (A : UU0) (f : coq_type M T - Proof. by move=> H; rewrite MS_bindE /cget H coerce_Some. Qed. Arguments Some_cget {T r} s. -(* Prove the laws *) -Definition cnewget T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A) : +Lemma cnewget T (s : coq_type M T) A (k : loc T -> coq_type M T -> M A) : cnew s >>= (fun r => cget r >>= k r) = cnew s >>= (fun r => k r s). Proof. -(*apply/boolp.funext => st/=. -rewrite bindE. -transitivity ( - Join - ((stateT (seq (binding idfun)) option_monad # (fun r : loc T => cget r >>= k r)) - (cnew s)) st -) => //.*) apply/boolp.funext => e. by rewrite bind_cnew (Some_cget s)// nth_error_rcons_size. Qed. -Definition cnewput T (s t : coq_type M T) A (k : loc T -> M A) : +Lemma cnewput T (s t : coq_type M T) A (k : loc T -> M A) : cnew s >>= (fun r => cput r t >> k r) = cnew t >>= k. Proof. apply/boolp.funext => e. @@ -1845,24 +1766,6 @@ rewrite bind_cnew 2!MS_bindE. by rewrite /cput/= nth_error_rcons_size coerce_Some set_nth_rcons. Qed. -Variant nth_error_spec (T : MLU) (e : Env) (r : loc T) : Type := - | NthError : forall s : coq_type M T, - nth_error e (loc_id r) = Some (mkbind T s) -> nth_error_spec e r - | NthError_nocoerce : forall T' (s' : coq_type M T'), - nth_error e (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> - nth_error_spec e r - | NthError_None : nth_error e (loc_id r) = None -> nth_error_spec e r. - -Lemma ntherrorP (T : MLU) (e : Env) (r : loc T) : nth_error_spec e r. -Proof. -case H : (nth_error e (loc_id r)) => [[T' s']|]. - have [Ts'|Ts'] := boolp.pselect (coerce T s'). - have ? := coerce_eq Ts'; subst T'. - exact: NthError H. - exact: NthError_nocoerce H Ts'. -exact: NthError_None. -Qed. - Lemma nocoerce_cget T (r : loc T) T' (s' : coq_type M T') e : nth_error e (loc_id r) = Some (mkbind T' s') -> ~ coerce T s' -> @@ -1904,7 +1807,7 @@ Proof. by move=> H; rewrite /cput/= H coerce_Some/= nth_error_set_nth_id. Qed. -Definition cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. +Lemma cgetputskip T (r : loc T) : cget r >>= cput r = cget r >> skip. Proof. apply/boolp.funext => e /=. have [s' H|T' s' H Ts'|H] := ntherrorP e r. @@ -1913,8 +1816,8 @@ have [s' H|T' s' H Ts'|H] := ntherrorP e r. - by rewrite !MS_bindE None_cget. Qed. -Definition cgetget T (r : loc T) (A : UU0) - (k : coq_type M T -> coq_type M T -> M A) : +Lemma cgetget T (r : loc T) (A : UU0) + (k : coq_type M T -> coq_type M T -> M A) : cget r >>= (fun s => cget r >>= k s) = cget r >>= fun s => k s s. Proof. apply/boolp.funext => e /=. @@ -1940,8 +1843,8 @@ by rewrite bindE/= (Some_cget s)// nth_error_set_nth. Qed. Arguments Some_cputget {T} s'. -Definition cputget T (r: loc T) (s: coq_type M T) (A: UU0) - (k: coq_type M T -> M A) : +Lemma cputget T (r : loc T) (s : coq_type M T) (A : UU0) + (k : coq_type M T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. apply/boolp.funext => e /=. @@ -1965,8 +1868,8 @@ rewrite nth_error_set_nth coerce_Some/=. by rewrite set_set_nth eqxx. Qed. -Definition cputput T (r : loc T) (s s' : coq_type M T) : - cput r s >> cput r s' = cput r s'. +Lemma cputput T (r : loc T) (s s' : coq_type M T) : + cput r s >> cput r s' = cput r s'. Proof. apply/boolp.funext => e. have [s'' H|T'' s'' H Ts''|H] := ntherrorP e r. @@ -1975,7 +1878,7 @@ have [s'' H|T'' s'' H Ts''|H] := ntherrorP e r. - by rewrite MS_bindE !None_cput. Qed. -Definition cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) +Lemma cgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (A : UU0) (k : coq_type M T1 -> coq_type M T2 -> M A) : cget r1 >>= (fun u => cget r2 >>= (fun v => k u v)) = cget r2 >>= (fun v => cget r1 >>= (fun u => k u v)). @@ -2009,7 +1912,7 @@ rewrite bind_cnew//. by rewrite (Some_cget v) // (nth_error_rcons_some _ H). Qed. -Definition cgetnewD T T' (r : loc T) (s : coq_type M T') A +Lemma cgetnewD T T' (r : loc T) (s : coq_type M T') A (k : loc T' -> coq_type M T -> coq_type M T -> M A) : cget r >>= (fun u => cnew s >>= fun r' => cget r >>= k r' u) = cget r >>= (fun u => cnew s >>= fun r' => k r' u u). @@ -2021,7 +1924,7 @@ have [u Hr|u T1 Hr T1u|Hr] := ntherrorP e r. - by rewrite !MS_bindE None_cget. Qed. -Definition cgetnewE T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) +Lemma cgetnewE T1 T2 (r1 : loc T1) (s : coq_type M T2) (A : UU0) (k1 k2 : loc T2 -> M A) : (forall r2 : loc T2, loc_id r1 != loc_id r2 -> k1 r2 = k2 r2) -> cget r1 >> (cnew s >>= k1) = cget r1 >> (cnew s >>= k2). @@ -2035,7 +1938,7 @@ have [u r1u|T' s' Hr1 T1s'|Hr] := ntherrorP e r1. - by rewrite !MS_bindE None_cget. Qed. -Definition cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : +Lemma cgetputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s : coq_type M T2) : cget r1 >> cput r2 s = cput r2 s >> cget r1 >> skip. Proof. apply/boolp.funext => e /=. @@ -2100,7 +2003,7 @@ have [v Hr2|T2' v Hr2 T2v|Hr2] := ntherrorP e r2; last first. by rewrite (@nth_error_set_nth_other _ _ _ _ _ (mkbind T1' s1')). Qed. -Definition cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) +Lemma cputC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) (s2 : coq_type M T2) (A : UU0) : loc_id r1 != loc_id r2 \/ JMeq s1 s2 -> cput r1 s1 >> cput r2 s2 = cput r2 s2 >> cput r1 s1. @@ -2149,7 +2052,7 @@ have [u Hr1|T1' s'd Hr1 T1s'|Hr1] := ntherrorP e r1; last first. by rewrite set_set_nth (negbTE Hr). Qed. -Definition cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) +Lemma cputgetC T1 T2 (r1 : loc T1) (r2 : loc T2) (s1 : coq_type M T1) (A : UU0) (k : coq_type M T2 -> M A) : loc_id r1 != loc_id r2 -> cput r1 s1 >> cget r2 >>= k = cget r2 >>= (fun v => cput r1 s1 >> k v). @@ -2188,7 +2091,7 @@ rewrite /bindS /= MS_mapE /= fmapE Hr1. by rewrite coerce_Some. Qed. -Definition cputnewC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A +Lemma cputnewC T T' (r : loc T) (s : coq_type M T) (s' : coq_type M T') A (k : loc T' -> M A) : cget r >> (cnew s' >>= fun r' => cput r s >> k r') = cput r s >> (cnew s' >>= k). @@ -2267,12 +2170,13 @@ Qed. HB.instance Definition _ := Monad.on M. HB.instance Definition isMonadTypedStoreModel := - isMonadTypedStore.Build MLU locT M cnewget cnewput cgetput cgetputskip + isMonadTypedStore.Build MLU locT_nat M cnewget cnewput cgetput cgetputskip cgetget cputget cputput cgetC cgetnewD cgetnewE cgetputC cputC cputgetC cputnewC crunret crunskip crunnew crunnewget crungetnew crungetput. -End typed_store. +End mkbind. +End ModelTypedStore. End ModelTypedStore. (* TODO? diff --git a/monad_transformer.v b/monad_transformer.v index c0e32b59..6e6c3ba2 100644 --- a/monad_transformer.v +++ b/monad_transformer.v @@ -149,6 +149,10 @@ HB.instance Definition _ := isMonadM_ret_bind.Build End state_monad_transformer. +Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M B) s : + (m >>= f) s = m s >>= uncurry f. +Proof. by []. Qed. + Definition stateT (S : UU0) := fun M => [the monad of MS S M]. HB.instance Definition _ (S : UU0) := isMonadT.Build diff --git a/monae_lib.v b/monae_lib.v index 81dc4417..8075f59a 100644 --- a/monae_lib.v +++ b/monae_lib.v @@ -6,16 +6,18 @@ Require ProofIrrelevance. Definition proof_irr := boolp.Prop_irrelevance. -Definition eq_rect_eq := @ProofIrrelevance.ProofIrrelevanceTheory.Eq_rect_eq.eq_rect_eq. +Definition eq_rect_eq := + @ProofIrrelevance.ProofIrrelevanceTheory.Eq_rect_eq.eq_rect_eq. Definition funext_dep := boolp.functional_extensionality_dep. (******************************************************************************) (* Shared notations and easy definitions/lemmas of general interest *) (* *) -(* foldr1 *) -(* curry/uncurry == currying for pairs *) -(* curry3/uncurry3 == currying for triples *) +(* foldr1 *) +(* curry/uncurry == currying for pairs *) +(* curry3/uncurry3 == currying for triples *) +(* coerce T1 (v : f T1) == some (f T2) if T1 = T2 and None o.w. *) (* *) (******************************************************************************) @@ -141,3 +143,90 @@ Proof. by []. Qed. Lemma if_pair A B b (x : A) y (u : A) (v : B) : (if b then (x, y) else (u, v)) = (if b then x else u, if b then y else v). Proof. by case: ifPn. Qed. + +Lemma iteriSr T n (f : nat -> T -> T) x : + iteri n.+1 f x = iteri n (f \o succn) (f 0 x). +Proof. by elim: n x => // n IH x /=; rewrite -IH. Qed. + +Lemma iteriD T n m (f : nat -> T -> T) x : + iteri (n + m) f x = iteri m (f \o addn n) (iteri n f x). +Proof. by elim: n x f => // n IH x f; rewrite addSn iteriSr IH iteriSr. Qed. + +Section nth_error. +Context [T : Type] (def : T) (st : seq T). + +(* Basic lemmas for standard library's nth_error *) +Local Notation nth_error := List.nth_error. + +Lemma nth_error_set_nth n x : nth_error (set_nth def st n x) n = Some x. +Proof. +elim: n st => [|z IH] [] // {IH}. +elim: z.+1 => [|n <-] //=. +by rewrite set_nth_nil. +Qed. + +Lemma nth_error_rcons_size b : nth_error (rcons st b) (size st) = Some b. +Proof. by elim: st. Qed. + +Lemma nth_error_rcons_some n a b : + nth_error st n = Some a -> nth_error (rcons st b) n = Some a. +Proof. by elim: n st => [|n IH] []. Qed. + +Lemma nth_error_set_nth_id n a : + nth_error st n = Some a -> set_nth def st n a = st. +Proof. by elim: n st => [|n IH] [] //= b st'; [case=> -> | move/IH ->]. Qed. + +Lemma nth_error_set_nth_other m n a b : + m != n -> + nth_error st m = Some a -> + nth_error (set_nth def st n b) m = Some a. +Proof. +elim: m st n => [|m IH] [|c st'] [|n] //=; rewrite eqSS => *; exact: IH. +Qed. + +Lemma nth_error_set_nth_none m n a b : + nth_error st m = None -> + nth_error st n = Some a -> + nth_error (set_nth def st n b) m = None. +Proof. by elim: m st n => [|m IH] [|c st'] [|n] //=; apply IH. Qed. + +Lemma nth_error_size n a : nth_error st n = Some a -> n < size st. +Proof. by elim: n st => [|n IH] [|c st'] //= /IH. Qed. + +Lemma nth_error_size_set_nth n a b : + nth_error st n = Some a -> size (set_nth def st n b) = size st. +Proof. by rewrite size_set_nth => /nth_error_size /maxn_idPr. Qed. + +Lemma set_nth_rcons a b : set_nth def (rcons st a) (size st) b = rcons st b. +Proof. by elim: st => //= c st' ->. Qed. + +Lemma nth_error_set_nth_rcons n a b c : + nth_error st n = Some a -> + set_nth def (rcons st c) n b = rcons (set_nth def st n b) c. +Proof. by elim: n st => [|n IH] [|d st'] //= /IH ->. Qed. +End nth_error. +Arguments nth_error_size {T st n a}. + +Section coerce. +Variables (X : eqType) (f : X -> Type). + +Definition coerce (T1 T2 : X) (v : f T1) : option (f T2) := + if @eqP _ T1 T2 is ReflectT H then Some (eq_rect _ _ v _ H) else None. + +Lemma coerce_sym (T T' : X) (s : f T) (s' : f T') : coerce T' s -> coerce T s'. +Proof. +by rewrite /coerce; case: eqP => //= h; case: eqP => //; rewrite h; auto. +Qed. + +Lemma coerce_Some (T : X) (s : f T) : coerce T s = Some s. +Proof. +by rewrite /coerce; case: eqP => /= [?|]; [rewrite -eq_rect_eq|auto]. +Qed. + +Lemma coerce_eq (T T' : X) (s : f T) : coerce T' s -> T = T'. +Proof. by rewrite /coerce; case: eqP. Qed. + +Lemma coerce_None (T T' : X) (s : f T) : T != T' -> coerce T' s = None. +Proof. by rewrite /coerce; case: eqP. Qed. + +End coerce. diff --git a/typed_store_lib.v b/typed_store_lib.v index 633e6411..4a3b2cbe 100644 --- a/typed_store_lib.v +++ b/typed_store_lib.v @@ -1,5 +1,5 @@ (* monae: Monadic equational reasoning in Coq *) -(* Copyright (C) 2020 monae authors, license: LGPL-2.1-or-later *) +(* Copyright (C) 2023 monae authors, license: LGPL-2.1-or-later *) Ltac typeof X := type of X. Require Import ssrmatching Reals JMeq. @@ -8,6 +8,10 @@ From mathcomp Require boolp. From infotheo Require Import Reals_ext. Require Import monae_lib hierarchy. +(******************************************************************************) +(* Lemmas using the typed store monad *) +(******************************************************************************) + Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. diff --git a/typed_store_model.v b/typed_store_model.v index 901ee0d1..fc6d738c 100644 --- a/typed_store_model.v +++ b/typed_store_model.v @@ -1,16 +1,21 @@ -From mathcomp Require Import all_ssreflect. -From mathcomp Require Import finmap. +(* monae: Monadic equational reasoning in Coq *) +(* Copyright (C) 2023 monae authors, license: LGPL-2.1-or-later *) +Require Import JMeq. +From mathcomp Require Import all_ssreflect finmap. From mathcomp Require boolp. #[local] Remove Hints boolp.Prop_irrelevance : core. Require Import monae_lib. From HB Require Import structures. Require Import hierarchy monad_lib fail_lib state_lib trace_lib. Require Import monad_transformer monad_model. -Require Import JMeq. (******************************************************************************) -(* Models for typed store *) -(* (separate file as it requires disabling various sanity checks) *) +(* Model for typed store *) +(* *) +(* Separate file as it requires disabling various sanity checks. *) +(* Reuses coerce and locT_nat from monad_model.v. *) +(* Similarities with ModelTypedStore from monad_model.v *) +(* *) (******************************************************************************) Set Implicit Arguments. @@ -31,7 +36,8 @@ Definition M0 Env (T : UU0) := MS Env option_monad T. End ModelTypedStore. #[bypass_check(positivity)] -Inductive Env (MLU : ML_universe) := mkEnv : seq (binding MLU (M0 (Env _))) -> Env _. +Inductive Env (MLU : ML_universe) := + mkEnv : seq (binding MLU (M0 (Env _))) -> Env _. Section ModelTypedStore_contd. Variable MLU : ML_universe. @@ -47,9 +53,9 @@ Local Notation M := acto. Local Notation coq_type := (@coq_type MLU M). -Local Notation undef := (val_nonempty MLU). +Local Notation val_nonundef := (val_nonempty MLU). -Definition def := mkbind (undef M). +Definition def := mkbind (val_nonempty MLU M). Local Notation nth_error := List.nth_error. @@ -57,7 +63,7 @@ Definition extend_env T (v : coq_type T) (e : Env) := mkEnv (rcons (ofEnv e) (mkbind v)). Definition fresh_loc (T : MLU) (e : Env) := mkloc T (sizeEnv e). -Local Notation loc := (@loc MLU locT). +Local Notation loc := (@loc MLU locT_nat). Definition cnew T (v : coq_type T) : M (loc T) := fun e => inr (fresh_loc T e, extend_env v e). @@ -84,11 +90,6 @@ Definition crun (A : UU0) (m : M A) : option A := | inr (a, _) => Some a end. -(* WIP *) -Lemma MS_bindE [S : UU0] [M : monad] [A B : UU0] (m : MS S M A) (f : A -> MS S M B) s : - (m >>= f) s = m s >>= uncurry f. -Proof. by []. Qed. - Lemma bind_cnew T (s : coq_type T) A (k : loc T -> M A) e : (cnew s >>= k) e = k (fresh_loc T e) (extend_env s e). Proof. by case: e. Qed. @@ -233,7 +234,7 @@ have [s' H|T' s' H Ts'|H] := ntherrorP e r. - by rewrite !MS_bindE None_cget. Qed. -Lemma cputget T (r: loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : +Lemma cputget T (r : loc T) (s: coq_type T) (A: UU0) (k: coq_type T -> M A) : cput r s >> (cget r >>= k) = cput r s >> k s. Proof. apply/boolp.funext => e /=. From 2aa98a9d8ddc21fcad800cc45afb2600f7585c59 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 25 Jul 2023 16:12:40 +0900 Subject: [PATCH 81/82] nitpick --- example_typed_store.v | 13 ++++++------- 1 file changed, 6 insertions(+), 7 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 999c33a5..12dc0062 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -219,11 +219,11 @@ pose y := x. rewrite -{2}/y. pose i := n. rewrite -[in LHS]/i. -have : (x, y) = (fibo_rec (n-i), fibo_rec (n.+1-i)). +have : (x, y) = (fibo_rec (n - i), fibo_rec (n.+1 - i)). by rewrite subnn -addn1 addKn. have : i <= n by []. elim: i x y => [|i IH] x y Hi. - rewrite !subn0 => -[] -> ->. + rewrite !subn0 => -[-> ->]. rewrite -/(fibo_rec n.+1). under eq_bind do under eq_bind do rewrite /= bindskipf. rewrite -cnewchk. @@ -231,11 +231,10 @@ elim: i x y => [|i IH] x y Hi. by rewrite cnewget -bindA crunret // crunnew // crunnew0. rewrite subSS => -[] Hx Hy. rewrite -(IH y (x + y) (ltnW Hi)); last first. - subst x y. - congr pair. - case: n Hi {IH} => // n. - rewrite subSS ltnS => Hi. - by rewrite -addn2 -addn1 -addnBAC // -addnBAC // addn2 addn1. + rewrite {}Hx {}Hy; congr pair. + rewrite subSn 1?ltnW//. + case: n {IH} => // n in Hi *. + by rewrite [in RHS]subSn -1?ltnS// subSS subSn -1?ltnS. rewrite /=. under eq_bind do under eq_bind do rewrite !bindA. rewrite -cnewchk. From a2cd43ea8bf29e362d5916ce3e9c9d0577859d5b Mon Sep 17 00:00:00 2001 From: Takafumi Saikawa Date: Tue, 25 Jul 2023 16:53:20 +0900 Subject: [PATCH 82/82] improve notations and --- example_typed_store.v | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/example_typed_store.v b/example_typed_store.v index 12dc0062..ae8a0514 100644 --- a/example_typed_store.v +++ b/example_typed_store.v @@ -54,7 +54,7 @@ Definition val_nonempty (M : UU0 -> UU0) := tt. Definition locT := [eqType of nat]. -Notation loc := (@loc ml_type locT). +Notation loc := (@loc _ locT). Inductive rlist (a : Type) (a_1 : ml_type) := | Nil @@ -94,7 +94,7 @@ Canonical MLU := @monad_model.MLU _ coq_type_nat0 ml_unit (val_nonempty idfun). Section cyclic. Variables (M : typedStoreMonad MLU MLTypes.locT). -Notation coq_type := (@hierarchy.coq_type MLU M). +Local Notation coq_type := (hierarchy.coq_type (MonadTypedStore.sort M)). Local Open Scope do_notation. Definition cycle (T : ml_type) (a b : coq_type T) @@ -155,7 +155,7 @@ End factorial. Section fact_for. Variable M : typedStoreMonad MLU MLTypes.locT. -Notation coq_type := (@coq_type MLU M). +Local Notation coq_type := (hierarchy.coq_type (MonadTypedStore.sort M)). Local Open Scope do_notation. Definition fact_for (n : coq_type ml_int) : M (coq_type ml_int) := @@ -196,7 +196,7 @@ End fact_for. Section fibonacci. Variable M : typedStoreMonad MLU MLTypes.locT. -Notation coq_type := (@coq_type M). +Local Notation coq_type := (hierarchy.coq_type (MonadTypedStore.sort M)). Fixpoint fibo_rec n := if n is m.+1 then @@ -425,7 +425,7 @@ Canonical ml_type63 := @Build_ML_universe _ coq_type63 ml_unit val_nonempty. Section fact_for_int63. Variable M : typedStoreMonad ml_type63 MLTypes.locT. -Notation coq_type := (@MLtypes63.coq_type63 M). +Local Notation coq_type := (hierarchy.coq_type (MonadTypedStore.sort M)). Local Open Scope do_notation. Section forloop63.