diff --git a/algebra/upred_tactics.v b/algebra/upred_tactics.v
index 0540bf34f381da70d418a3c14e2db24c95178a4d..4e82798c749fe424ff4fafe55cc7e629f8821a0f 100644
--- a/algebra/upred_tactics.v
+++ b/algebra/upred_tactics.v
@@ -151,19 +151,19 @@ Tactic Notation "ecancel" open_constr(Ps) :=
     Will turn this goal into P ⊑ Q and strip ▷ in P below ★, ∧, ∨. *)
 Ltac strip_later :=
   let rec strip :=
-      lazymatch goal with
-      | |- (_ ★ _) ⊑ ▷ _  =>
-        etrans; last (eapply equiv_entails_sym, later_sep);
-        apply sep_mono; strip
-      | |- (_ ∧ _) ⊑ ▷ _  =>
-        etrans; last (eapply equiv_entails_sym, later_and);
-        apply sep_mono; strip
-      | |- (_ ∨ _) ⊑ ▷ _  =>
-        etrans; last (eapply equiv_entails_sym, later_or);
-        apply sep_mono; strip
-      | |- ▷ _ ⊑ ▷ _ => apply later_mono; reflexivity
-      | |- _ ⊑ ▷ _ => apply later_intro; reflexivity
-      end
+    lazymatch goal with
+    | |- (_ ★ _) ⊑ ▷ _  =>
+      etrans; last (eapply equiv_entails_sym, later_sep);
+      apply sep_mono; strip
+    | |- (_ ∧ _) ⊑ ▷ _  =>
+      etrans; last (eapply equiv_entails_sym, later_and);
+      apply sep_mono; strip
+    | |- (_ ∨ _) ⊑ ▷ _  =>
+      etrans; last (eapply equiv_entails_sym, later_or);
+      apply sep_mono; strip
+    | |- ▷ _ ⊑ ▷ _ => apply later_mono; reflexivity
+    | |- _ ⊑ ▷ _ => apply later_intro; reflexivity
+    end
   in let rec shape_Q :=
     lazymatch goal with
     | |- _ ⊑ (_ ★ _) =>
@@ -190,13 +190,14 @@ Ltac strip_later :=
 (* TODO: this name may be a big too general *)
 Ltac revert_all :=
   lazymatch goal with
-  | |- ∀ _, _ => let H := fresh in intro H; revert_all;
-           (* TODO: Really, we should distinguish based on whether this is a
-              dependent function type or not. Right now, we distinguish based
-              on the sort of the argument, which is suboptimal. *)
-           first [ apply (const_intro_impl _ _ _ H); clear H
-                 | revert H; apply forall_elim']
-  | |- ?C ⊑ _ => apply impl_entails
+  | |- ∀ _, _ =>
+    let H := fresh in intro H; revert_all;
+    (* TODO: Really, we should distinguish based on whether this is a
+    dependent function type or not. Right now, we distinguish based
+    on the sort of the argument, which is suboptimal. *)
+    first [ apply (const_intro_impl _ _ _ H); clear H
+          | revert H; apply forall_elim']
+  | |- _ ⊑ _ => apply impl_entails
   end.
 
 (** This starts on a goal of the form ∀ ..., ?0... → ?1 ⊑ ?2.
@@ -217,16 +218,15 @@ Ltac löb tac :=
   (* Now introduce again all the things that we reverted, and at the bottom,
      do the work *)
   let rec go :=
-      lazymatch goal with
-      | |- _ ⊑ (∀ _, _) => apply forall_intro;
-               let H := fresh in intro H; go; revert H
-      | |- _ ⊑ (■ _ → _) => apply impl_intro_l, const_elim_l;
-               let H := fresh in intro H; go; revert H
-      (* This is the "bottom" of the goal, where we see the impl introduced
-         by uPred_revert_all as well as the ▷ from löb_strong and the □ we added. *)
-      | |- ▷ □ ?R ⊑ (?L → _) => apply impl_intro_l;
-               trans (L ★ ▷ □ R)%I;
-                 first (eapply equiv_entails, always_and_sep_r, _; reflexivity);
-               tac
-      end
+    lazymatch goal with
+    | |- _ ⊑ (∀ _, _) =>
+      apply forall_intro; let H := fresh in intro H; go; revert H
+    | |- _ ⊑ (■ _ → _) =>
+      apply impl_intro_l, const_elim_l; let H := fresh in intro H; go; revert H
+    (* This is the "bottom" of the goal, where we see the impl introduced
+       by uPred_revert_all as well as the ▷ from löb_strong and the □ we added. *)
+    | |- ▷ □ ?R ⊑ (?L → _) => apply impl_intro_l;
+      trans (L ★ ▷ □ R)%I;
+        [eapply equiv_entails, always_and_sep_r, _; reflexivity | tac]
+    end
   in go.
diff --git a/heap_lang/wp_tactics.v b/heap_lang/wp_tactics.v
index 0f5f385486f9504877593c1ab9e5d08157b43738..1eb190fb3bb218a635b6e1158e32673d84041ce7 100644
--- a/heap_lang/wp_tactics.v
+++ b/heap_lang/wp_tactics.v
@@ -15,25 +15,25 @@ Ltac wp_finish :=
   match goal with
   | |- _ ⊑ ▷ _ => etrans; [|apply later_mono; go; reflexivity]
   | |- _ ⊑ wp _ _ _ =>
-     etrans; [|eapply wp_value_pvs; reflexivity];
-     (* sometimes, we will have to do a final view shift, so only apply
-     pvs_intro if we obtain a consecutive wp *)
-     try (eapply pvs_intro;
-          match goal with |- _ ⊑ wp _ _ _ => simpl | _ => fail end)
+    etrans; [|eapply wp_value_pvs; reflexivity];
+    (* sometimes, we will have to do a final view shift, so only apply
+    pvs_intro if we obtain a consecutive wp *)
+    try (eapply pvs_intro;
+         match goal with |- _ ⊑ wp _ _ _ => simpl | _ => fail end)
   | _ => idtac
   end in simpl; intros_revert go.
 
 Tactic Notation "wp_rec" ">" :=
-  löb ltac:((* Find the redex and apply wp_rec *)
-              idtac; (* <https://coq.inria.fr/bugs/show_bug.cgi?id=4584> *)
-               lazymatch goal with
-               | |- _ ⊑ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
-                        match eval cbv in e' with
-                        | App (Rec _ _ _) _ =>
-                          wp_bind K; etrans; [|eapply wp_rec; reflexivity];
-                          wp_finish
-                        end)
-               end).
+  löb ltac:(
+    (* Find the redex and apply wp_rec *)
+    idtac; (* <https://coq.inria.fr/bugs/show_bug.cgi?id=4584> *)
+    lazymatch goal with
+    | |- _ ⊑ wp ?E ?e ?Q => reshape_expr e ltac:(fun K e' =>
+      match eval cbv in e' with
+      | App (Rec _ _ _) _ =>
+         wp_bind K; etrans; [|eapply wp_rec; reflexivity]; wp_finish
+      end)
+     end).
 Tactic Notation "wp_rec" := wp_rec>; try strip_later.
 
 Tactic Notation "wp_lam" ">" :=