Skip to content
Snippets Groups Projects
proofmode.ref 27.7 KiB
Newer Older
"demo_0"
     : string
  BiPersistentlyForall0 : BiPersistentlyForall PROP
  P, Q : PROP
  ============================
  "H2" : ∀ x : nat, ⌜x = 0⌝ ∨ ⌜x = 1⌝
  --------------------------------------□
  "H" : □ (P ∨ Q)
  --------------------------------------∗
  Q ∨ P
  BiPersistentlyForall0 : BiPersistentlyForall PROP
  P, Q : PROP
  ============================
  "H2" : ∀ x : nat, ⌜x = 0⌝ ∨ ⌜x = 1⌝
  _ : P
  --------------------------------------□
  Q ∨ P
Ralf Jung's avatar
Ralf Jung committed
"test_iStopProof"
     : string
Ralf Jung's avatar
Ralf Jung committed
  Q : PROP
  ============================
  "H1" : emp
  --------------------------------------□
  "H2" : Q
  --------------------------------------∗
  Q
Ralf Jung's avatar
Ralf Jung committed
  Q : PROP
  ============================
"test_iDestruct_and_emp"
     : string
  P, Q : PROP
  Persistent0 : Persistent P
  Persistent1 : Persistent Q
  ============================
  _ : P
  _ : Q
  --------------------------------------□
  <affine> (P ∗ Q)
"test_iDestruct_spatial"
     : string
  Q : PROP
  ============================
  "HQ" : <affine> Q
  --------------------------------------∗
  Q
"test_iDestruct_spatial_affine"
     : string
  Q : PROP
  Affine0 : Affine Q
  ============================
  "HQ" : Q
  --------------------------------------∗
  Q
"test_iDestruct_exists_not_exists"
     : string
The command has indeed failed with message:
Tactic failure: iExistDestruct: cannot destruct P.
"test_iDestruct_exists_intuitionistic"
     : string
  
  PROP : bi
  P : PROP
  Φ : nat → PROP
  y : nat
  ============================
  "H" : Φ y ∧ P
  --------------------------------------□
  P
"test_iDestruct_exists_anonymous"
     : string
  
  PROP : bi
  P : PROP
  Φ : nat → PROP
  H : nat
  ============================
  "HΦ" : ∃ x : nat, Φ x
  --------------------------------------∗
  ∃ x : nat, Φ x
"test_iDestruct_nameless_exist"
     : string
1 goal
  
  PROP : bi
  Φ : nat → PROP
  __unknown : nat
  ============================
  "H" : Φ __unknown
  --------------------------------------∗
  ∃ x : nat, Φ x
"test_iIntros_nameless_forall"
     : string
1 goal
  
  PROP : bi
  Φ : nat → PROP
  __unknown : nat
  ============================
  "H" : ∀ x : nat, Φ x
  --------------------------------------∗
  Φ __unknown
"test_iIntros_nameless_pure_forall"
     : string
1 goal
  
  PROP : bi
  BiPureForall0 : BiPureForall PROP
  φ : nat → Prop
  __unknown : nat
  ============================
  "H" : ∀ x : nat, ⌜φ x⌝
  --------------------------------------∗
  ⌜φ __unknown⌝
"test_iIntros_forall_pure"
     : string
  Ψ : nat → PROP
  ============================
  --------------------------------------∗
  Ψ x → Ψ x
"test_iIntros_pure_names"
     : string
  
  PROP : bi
  H : True
  P : PROP
  x, y : nat
  H0 : x = y
  ============================
  --------------------------------------∗
  P -∗ P
The command has indeed failed with message:
No applicable tactic.
The command has indeed failed with message:
Tactic failure: iElaborateSelPat: "HQ" not found.
The command has indeed failed with message:
Tactic failure: iElaborateSelPat: "HQ" not found.
"test_iSpecialize_pure_error"
     : string
The command has indeed failed with message:
Tactic failure: iSpecialize: P not pure.
"test_iSpecialize_pure_error"
     : string
The command has indeed failed with message:
Tactic failure: iSpecialize: cannot solve φ using done.
"test_iSpecialize_done_error"
     : string
The command has indeed failed with message:
Tactic failure: iSpecialize: cannot solve P using done.
The command has indeed failed with message:
Tactic failure: iSpecialize: Q not persistent.
"test_iSpecialize_impl_pure"
     : string
1 goal
  
  PROP : bi
  φ : Prop
  P, Q : PROP
  H : φ
  ============================
  --------------------------------------∗
  <affine> ⌜φ⌝
1 goal
  
  PROP : bi
  φ : Prop
  P, Q : PROP
  H : φ
  ============================
  "H1" : P
  --------------------------------------□
  <affine> ⌜φ⌝
"test_iSpecialize_impl_pure_affine"
     : string
1 goal
  
  PROP : bi
  BiAffine0 : BiAffine PROP
  φ : Prop
  P, Q : PROP
  H : φ
  ============================
  --------------------------------------∗
  ⌜φ⌝
1 goal
  
  PROP : bi
  BiAffine0 : BiAffine PROP
  φ : Prop
  P, Q : PROP
  H : φ
  ============================
  "H1" : P
  --------------------------------------□
  ⌜φ⌝
"test_iSpecialize_impl_pure"
     : string
1 goal
  
  PROP : bi
  φ : Prop
  P, Q : PROP
  H : φ
  ============================
  --------------------------------------∗
  <affine> ⌜φ⌝
1 goal
  
  PROP : bi
  φ : Prop
  P, Q : PROP
  H : φ
  ============================
  "H1" : P
  --------------------------------------□
  <affine> ⌜φ⌝
"test_iSpecialize_impl_pure_affine"
     : string
1 goal
  
  PROP : bi
  BiAffine0 : BiAffine PROP
  φ : Prop
  P, Q : PROP
  H : φ
  ============================
  --------------------------------------∗
  ⌜φ⌝
1 goal
  
  PROP : bi
  BiAffine0 : BiAffine PROP
  φ : Prop
  P, Q : PROP
  H : φ
  ============================
  "H1" : P
  --------------------------------------□
"test_iAssert_intuitionistic"
     : string
The command has indeed failed with message:
Tactic failure: iSpecialize: (|==> P)%I not persistent.
The command has indeed failed with message:
Tactic failure: iFrame: cannot frame (Φ 0 1).
"test_iFrame_conjunction_3"
     : string
1 goal
  
  PROP : bi
  P, Q : PROP
  Absorbing0 : Absorbing Q
  ============================
  "HQ" : Q
  --------------------------------------∗
  False
Robbert Krebbers's avatar
Robbert Krebbers committed
"test_iFrame_affinely_emp"
     : string
1 goal
  
  PROP : bi
  P : PROP
  ============================
  "H" : P
  --------------------------------------□
  ∃ _ : nat, emp
"test_iFrame_affinely_True"
     : string
1 goal
  
  PROP : bi
  BiAffine0 : BiAffine PROP
  P : PROP
  ============================
  "H" : P
  --------------------------------------□
  ∃ _ : nat, True
"test_iFrame_or_1"
     : string
1 goal
  
  PROP : bi
  P1, P2, P3 : PROP
  ============================
  --------------------------------------∗
  ▷ (∃ _ : nat, emp)
"test_iFrame_or_2"
     : string
1 goal
  
  PROP : bi
  P1, P2, P3 : PROP
  ============================
  --------------------------------------∗
  ▷ (∃ x : nat, emp ∧ ⌜x = 0⌝ ∨ emp)
"test_iFrame_or_3"
     : string
1 goal
  
  PROP : bi
  P1, P2, P3 : PROP
  ============================
  --------------------------------------∗
  ▷ (∃ x : nat, ⌜x = 0⌝)
Robbert Krebbers's avatar
Robbert Krebbers committed
"test_iFrame_or_affine_1"
     : string
1 goal
  
  PROP : bi
  BiAffine0 : BiAffine PROP
  P1, P2, P3 : PROP
  ============================
  --------------------------------------∗
  ▷ (∃ _ : nat, True)
"test_iFrame_or_affine_2"
     : string
1 goal
  
  PROP : bi
  BiAffine0 : BiAffine PROP
  P1, P2, P3 : PROP
  ============================
  --------------------------------------∗
  ▷ (∃ _ : nat, True)
"test_iCombine_nested_no_gives"
     : string
The command has indeed failed with message:
Tactic failure: iCombine: cannot find 'gives' clause for hypotheses
["HP"; "HQ"].
The command has indeed failed with message:
Tactic failure: iCombine: cannot find 'gives' clause for hypotheses
["HP"; "HQ"].
"test_iInduction_multiple_IHs"
     : string
1 goal
  
  PROP : bi
  l, t1 : tree
  Φ : tree → PROP
  ============================
  "Hleaf" : Φ leaf
  "Hnode" : ∀ l0 r : tree, Φ l0 -∗ Φ r -∗ Φ (node l0 r)
  "IH" : Φ l
  "IH1" : Φ t1
  --------------------------------------□
  Φ (node l t1)
The command has indeed failed with message:
Tactic failure: iSpecialize: cannot instantiate (∀ _ : φ, P -∗ False)%I with
P.
The command has indeed failed with message:
Tactic failure: iSpecialize: cannot instantiate (⌜φ⌝ → P -∗ False)%I with P.
"test_iNext_plus_3"
     : string
  P, Q : PROP
  n, m, k : nat
  ============================
  --------------------------------------∗
  ▷^(S n + S m) emp
"test_specialize_nested_intuitionistic"
     : string
  φ : Prop
  P, P2, Q, R1, R2 : PROP
  H : φ
  ============================
  "HP" : P
  "HQ" : P -∗ Q
  --------------------------------------□
  "HR" : R2
  --------------------------------------∗
  R2
"test_iSimpl_in"
     : string
  x, y : nat
  ============================
  "H" : ⌜S (S (S x)) = y⌝
  --------------------------------------∗
  ⌜S (S (S x)) = y⌝
  x, y, z : nat
  ============================
  "H1" : ⌜S (S (S x)) = y⌝
  "H2" : ⌜S y = z⌝
  --------------------------------------∗
  ⌜S (S (S x)) = y⌝
  x, y, z : nat
  ============================
  "H1" : ⌜S (S (S x)) = y⌝
  --------------------------------------□
  "H2" : ⌜(1 + y)%nat = z⌝
  --------------------------------------∗
  ⌜S (S (S x)) = y⌝
"test_iSimpl_in4"
     : string
The command has indeed failed with message:
Tactic failure: iEval: %: unsupported selection pattern.
"test_iRename"
     : string
1 goal
  
  PROP : bi
  P : PROP
  ============================
  "X" : P
  --------------------------------------□
  P
"test_iFrame_later_1"
     : string
  P, Q : PROP
  ============================
  --------------------------------------∗
  ▷ emp
"test_iFrame_later_2"
     : string
  P, Q : PROP
  ============================
  --------------------------------------∗
  ▷ emp
The command has indeed failed with message:
Tactic failure: iFrame: cannot frame Q.
The command has indeed failed with message:
No applicable tactic.
The command has indeed failed with message:
No applicable tactic.
"test_and_sep_affine_bi"
     : string
  BiAffine0 : BiAffine PROP
  P, Q : PROP
  ============================
  _ : □ P
  _ : Q
  --------------------------------------∗
  □ P
"test_big_sepL_simpl"
     : string
  x : nat
  l : list nat
  P : PROP
  ============================
  "HP" : P
  _ : [∗ list] y ∈ l, <affine> ⌜y = y⌝
  _ : [∗ list] y ∈ (x :: l), <affine> ⌜y = y⌝
  --------------------------------------∗
  P
  x : nat
  l : list nat
  P : PROP
  ============================
  "HP" : P
  _ : [∗ list] y ∈ l, <affine> ⌜y = y⌝
  _ : <affine> ⌜x = x⌝ ∗ ([∗ list] y ∈ l, <affine> ⌜y = y⌝)
  --------------------------------------∗
  P
"test_big_sepL2_simpl"
     : string
  x1, x2 : nat
  l1, l2 : list nat
  P : PROP
  ============================
  "HP" : P
  _ : [∗ list] y1;y2 ∈ [];l2, <affine> ⌜y1 = y2⌝
  _ : [∗ list] y1;y2 ∈ (x1 :: l1);((x2 :: l2) ++ l2), <affine> ⌜y1 = y2⌝
  --------------------------------------∗
  P ∨ ([∗ list] _;_ ∈ (x1 :: l1);(x2 :: l2), True)
  x1, x2 : nat
  l1, l2 : list nat
  P : PROP
  ============================
  "HP" : P
  _ : [∗ list] y1;y2 ∈ [];l2, <affine> ⌜y1 = y2⌝
  _ : <affine> ⌜x1 = x2⌝ ∗
      ([∗ list] y1;y2 ∈ l1;(l2 ++ l2), <affine> ⌜y1 = y2⌝)
  --------------------------------------∗
  P ∨ True ∗ ([∗ list] _;_ ∈ l1;l2, True)
"test_big_sepL2_iDestruct"
     : string
  Φ : nat → nat → PROP
  x1, x2 : nat
  l1, l2 : list nat
  ============================
  _ : Φ x1 x2
  _ : [∗ list] y1;y2 ∈ l1;l2, Φ y1 y2
  --------------------------------------∗
  <absorb> Φ x1 x2
"test_reducing_after_iDestruct"
     : string
  ============================
  "H" : □ True
  --------------------------------------∗
  True
"test_reducing_after_iApply"
     : string
  ============================
  "H" : emp
  --------------------------------------□
  □ emp
"test_reducing_after_iApply_late_evar"
     : string
  ============================
  "H" : emp
  --------------------------------------□
  □ emp
"test_wandM"
     : string
  mP : option PROP
  Q, R : PROP
  ============================
  "HPQ" : mP -∗? Q
  "HQR" : Q -∗ R
  "HP" : default emp mP
  --------------------------------------∗
  R
  mP : option PROP
  Q, R : PROP
  ============================
  "HP" : default emp mP
  --------------------------------------∗
"test_iApply_prettification3"
     : string
1 goal
  
  PROP : bi
  Ψ, Φ : nat → PROP
  HP : ∀ (f : nat → nat) (y : nat),
         TCEq f (λ x : nat, x + 10) → Ψ (f 1) -∗ Φ y
  ============================
  "H" : Ψ 11
  --------------------------------------∗
  Ψ (1 + 10)
  P, Q : PROP
  ============================
  "HP" : ▷ P
  "HQ" : Q
  --------------------------------------∗
  ▷ P ∗ Q
Robbert Krebbers's avatar
Robbert Krebbers committed
"test_iRevert_pure"
     : string
1 goal
  
  PROP : bi
  φ : Prop
  P : PROP
  ============================
  "H" : <affine> ⌜φ⌝ -∗ P
  --------------------------------------∗
  <affine> ⌜φ⌝ -∗ P
Robbert Krebbers's avatar
Robbert Krebbers committed
"test_iRevert_order_and_names"
     : string
1 goal
  
  PROP : bi
  ============================
  --------------------------------------∗
  ∀ P1 P2, P1 -∗ P2 -∗ P1 ∗ P2
Robbert Krebbers's avatar
Robbert Krebbers committed
"test_iRevert_pure_affine"
     : string
1 goal
  
  PROP : bi
  BiAffine0 : BiAffine PROP
  φ : Prop
  P : PROP
  ============================
  "H" : ⌜φ⌝ -∗ P
  --------------------------------------∗
  ⌜φ⌝ -∗ P
"test_iFrame_not_add_emp_for_intuitionistically"
     : string
1 goal
  
  PROP : bi
  BiAffine0 : BiAffine PROP
  P : PROP
  ============================
  "H" : P
  --------------------------------------□
  ∃ _ : nat, True
Ralf Jung's avatar
Ralf Jung committed
"test_iIntros_auto_name_used_later"
     : string
The command has indeed failed with message:
x is already used.
"elim_mod_accessor"
     : string
  BiFUpd0 : BiFUpd PROP
  X : Type
  E1, E2 : coPset.coPset
  γ : X → option PROP
  ============================
  "Hacc" : ∃ x : X, α x ∗ (β x ={E2,E1}=∗ default emp (γ x))
  --------------------------------------∗
  |={E2,E1}=> True
"print_long_line_1"
     : string
  BiFUpd0 : BiFUpd PROP
  P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P : PROP
  ============================
  "HP" : P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P ∗
         P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P
  --------------------------------------∗
  BiFUpd0 : BiFUpd PROP
  P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P : PROP
  ============================
  _ : P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P ∗
      P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P
  --------------------------------------∗
"print_long_line_2"
     : string
  BiFUpd0 : BiFUpd PROP
  P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P : PROP
  ============================
  "HP" : TESTNOTATION {{ P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P |
         P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P }}
  --------------------------------------∗
  BiFUpd0 : BiFUpd PROP
  P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P : PROP
  ============================
  _ : TESTNOTATION {{ P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P |
      P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P_P }}
  --------------------------------------∗
"long_impl"
     : string
  BiFUpd0 : BiFUpd PROP
  PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
  ============================
  --------------------------------------∗
  PPPPPPPPPPPPPPPPP → QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗
  QQQQQQQQQQQQQQQQQQ
"long_impl_nested"
     : string
  BiFUpd0 : BiFUpd PROP
  PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
  ============================
  --------------------------------------∗
  PPPPPPPPPPPPPPPPP
  → QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
    → QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
"long_wand"
     : string
  BiFUpd0 : BiFUpd PROP
  PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
  ============================
  --------------------------------------∗
  PPPPPPPPPPPPPPPPP -∗
  QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
"long_wand_nested"
     : string
  BiFUpd0 : BiFUpd PROP
  PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
  ============================
  --------------------------------------∗
  PPPPPPPPPPPPPPPPP -∗
  QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ -∗
  QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
"long_fupd"
     : string
  BiFUpd0 : BiFUpd PROP
  E : coPset.coPset
  PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
  ============================
  --------------------------------------∗
  PPPPPPPPPPPPPPPPP ={E}=∗
  QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
"long_fupd_nested"
     : string
  BiFUpd0 : BiFUpd PROP
  E1, E2 : coPset.coPset
  PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
  ============================
  --------------------------------------∗
  PPPPPPPPPPPPPPPPP ={E1,E2}=∗
  QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ={E1,E2}=∗
  QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
"iStopProof_not_proofmode"
     : string
The command has indeed failed with message:
Tactic failure: iStopProof: proofmode not started.
"iAlways_spatial_non_empty"
     : string
The command has indeed failed with message:
Tactic failure: iModIntro: spatial context is non-empty.
"iDestruct_bad_name"
     : string
The command has indeed failed with message:
Tactic failure: iRename: "HQ" not found.
"iIntros_dup_name"
     : string
The command has indeed failed with message:
Tactic failure: iIntro: "HP" not fresh.
The command has indeed failed with message:
x is already used.
"iIntros_pure_not_forall"
     : string
The command has indeed failed with message:
Tactic failure: iIntro: cannot turn (P -∗ Q)%I into a universal quantifier.
The command has indeed failed with message:
Tactic failure: iSplitL: hypotheses ["HPx"] not found.
The command has indeed failed with message:
Tactic failure: iSplitL: hypotheses ["HPx"] not found.
"iSplitR_one_of_many"
     : string
The command has indeed failed with message:
Tactic failure: iSplitR: hypotheses ["HPx"] not found.
The command has indeed failed with message:
Tactic failure: iSplitR: hypotheses ["HPx"] not found.
"iSplitL_non_splittable"
     : string
The command has indeed failed with message:
Tactic failure: iSplitL: P not a separating conjunction.
"iSplitR_non_splittable"
     : string
The command has indeed failed with message:
Tactic failure: iSplitR: P not a separating conjunction.
"iCombine_fail"
     : string
The command has indeed failed with message:
Tactic failure: iCombine: hypotheses ["HP3"] not found.
"iSpecialize_bad_name1_fail"
     : string
The command has indeed failed with message:
Tactic failure: iSpecialize: "H" not found.
"iSpecialize_bad_name2_fail"
     : string
The command has indeed failed with message:
Tactic failure: iSpecialize: "H" not found.
Robbert Krebbers's avatar
Robbert Krebbers committed
"iSpecialize_autoframe_fail"
     : string
The command has indeed failed with message:
Tactic failure: iSpecialize: premise P cannot be solved by framing.
Robbert Krebbers's avatar
Robbert Krebbers committed
The command has indeed failed with message:
Tactic failure: iSpecialize: premise P cannot be solved by framing.
"iSpecialize_autoframe_fail2"
     : string
The command has indeed failed with message:
Tactic failure: iSpecialize: premise Q cannot be solved by framing.
The command has indeed failed with message:
Tactic failure: iSpecialize: premise Q cannot be solved by framing.
"iExact_fail"
     : string
The command has indeed failed with message:
Tactic failure: iExact: "HQ" not found.
The command has indeed failed with message:
Tactic failure: iExact: "HQ" : Q does not match goal.
The command has indeed failed with message:
Tactic failure: iExact: remaining hypotheses not affine and the goal not absorbing.
"iClear_fail"
     : string
The command has indeed failed with message:
Tactic failure: iElaborateSelPat: "HP" not found.
The command has indeed failed with message:
Tactic failure: iClear: "HP" : P not affine and the goal not absorbing.
"iSpecializeArgs_fail"
     : string
The command has indeed failed with message:
In environment
P : PROP
The term "true" has type "bool" while it is expected to have type "nat".
"iStartProof_fail"
     : string
The command has indeed failed with message:
Tactic failure: iStartProof: not a BI assertion: (0 = 0).
"iPoseProof_fail"
     : string
The command has indeed failed with message:
Tactic failure: iPoseProof: (0 = 0) not a BI assertion.
The command has indeed failed with message:
Tactic failure: iRename: "H" not fresh.
"iPoseProof_not_found_fail"
     : string
The command has indeed failed with message:
Tactic failure: iPoseProof: "Hx" not found.
"iPoseProof_not_found_fail2"
     : string
The command has indeed failed with message:
Tactic failure: iSpecialize: hypotheses ["HQ"] not found.
"iPoseProofCoreHyp_not_found_fail"
     : string
The command has indeed failed with message:
Tactic failure: iPoseProof: "Hx" not found.
"iPoseProofCoreHyp_not_fresh_fail"
     : string
The command has indeed failed with message:
Tactic failure: iPoseProof: "H1" not fresh.
"iRevert_fail"
     : string
The command has indeed failed with message:
Tactic failure: iElaborateSelPat: "H" not found.
"iDestruct_fail"
     : string
The command has indeed failed with message:
Tactic failure: iDestruct: "[{HP}]" has just a single conjunct.
The command has indeed failed with message:
Tactic failure: iDestruct: "// H" is not supported due to IDone.
The command has indeed failed with message:
Tactic failure: iDestruct: "HP //"
should contain exactly one proper introduction pattern.
The command has indeed failed with message:
Tactic failure: iDestruct: "[HP HQ HR]" has too many conjuncts.
The command has indeed failed with message:
Tactic failure: iDestruct: "[HP|HQ|HR]" has too many disjuncts.
The command has indeed failed with message:
Tactic failure: iDestruct: "[HP]" has just a single conjunct.
The command has indeed failed with message:
Tactic failure: iDestruct: in "[HP HQ|HR]" a disjunct has multiple patterns.
"iOrDestruct_fail"
     : string
The command has indeed failed with message:
Tactic failure: iOrDestruct: "H'" or "H2" not fresh.
The command has indeed failed with message:
Tactic failure: iOrDestruct: "H1" or "H'" not fresh.
"iApply_fail"
     : string
The command has indeed failed with message:
Tactic failure: iApply: cannot apply P.
"iApply_fail_not_affine_1"
     : string
The command has indeed failed with message:
Tactic failure: iApply: remaining hypotheses not affine and the goal not absorbing.
"iIntros_fail_nonempty_spatial"
     : string
The command has indeed failed with message:
Tactic failure: iIntro: introducing non-persistent "HP" : P
into non-empty spatial context.
"iIntros_fail_not_fresh"
     : string
The command has indeed failed with message:
Tactic failure: iIntro: "HP" not fresh.
"iIntros_fail_nothing_to_introduce"
     : string
The command has indeed failed with message:
Tactic failure: iIntro: could not introduce "HQ"
, goal is not a wand or implication.
"iApply_fail_not_affine_2"
     : string
The command has indeed failed with message:
Tactic failure: iApply: remaining hypotheses not affine and the goal not absorbing.
"iAssumption_fail_not_affine_1"
     : string
The command has indeed failed with message:
Tactic failure: iAssumption: remaining hypotheses not affine and the goal not absorbing.
"iAssumption_fail_not_affine_2"
     : string
The command has indeed failed with message:
Tactic failure: iAssumption: remaining hypotheses not affine and the goal not absorbing.
"iRevert_wrong_var"
     : string
The command has indeed failed with message:
Tactic failure: iRevert: k1 not in scope.
The command has indeed failed with message:
Tactic failure: iRevert: k1 not in scope.
"iRevert_dup_var"
     : string
The command has indeed failed with message:
Tactic failure: iRevert: k not in scope.
"iRevert_dep_var_coq"
     : string
The command has indeed failed with message:
k is used in hypothesis Hk.
"iRevert_dep_var"
     : string
The command has indeed failed with message:
Tactic failure: iRevert: k is used in hypothesis "Hk".
Robbert Krebbers's avatar
Robbert Krebbers committed
     : string
The command has indeed failed with message:
Tactic failure: iLöb: no 'BiLöb' instance found.
"iMod_mask_mismatch"
     : string
The command has indeed failed with message:
Tactic failure: 
"Goal and eliminated modality must have the same mask.
Use [iMod (fupd_mask_subseteq E2)] to adjust the mask of your goal to [E2]".
The command has indeed failed with message:
Tactic failure: