Newer
Older
1 subgoal
PROP : sbi
P, Q : PROP
============================
"H2" : ∀ x : nat, ⌜x = 0⌝ ∨ ⌜x = 1⌝
--------------------------------------□
"H" : □ (P ∨ Q)
--------------------------------------∗
Q ∨ P
1 subgoal
PROP : sbi
P, Q : PROP
============================
"H2" : ∀ x : nat, ⌜x = 0⌝ ∨ ⌜x = 1⌝
_ : P
--------------------------------------□
Q ∨ P
P, Q : PROP
Persistent0 : Persistent P
Persistent1 : Persistent Q
============================
_ : P
_ : Q
--------------------------------------□
<affine> (P ∗ Q)
The command has indeed failed with message:
Ltac call to "done" failed.
No applicable tactic.
The command has indeed failed with message:
In nested Ltac calls to "iClear (constr)", "iElaborateSelPat" and
"<iris.proofmode.ltac_tactics.iElaborateSelPat_go>", last call failed.
Tactic failure: iElaborateSelPat: "HQ" not found.
The command has indeed failed with message:
In nested Ltac calls to "iClear (constr)", "iElaborateSelPat" and
"<iris.proofmode.ltac_tactics.iElaborateSelPat_go>", last call failed.
Tactic failure: iElaborateSelPat: "HQ" not found.
The command has indeed failed with message:
In nested Ltac calls to "iSpecialize (open_constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) with (open_constr) (open_constr) as (constr)",
"iSpecializePat (open_constr) (constr)" and "iSpecializePat_go", last call
failed.
Tactic failure: iSpecialize: cannot instantiate (∀ _ : φ, P -∗ False)%I with
P.
The command has indeed failed with message:
In nested Ltac calls to "iSpecialize (open_constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) with (open_constr) (open_constr) as (constr)",
"iSpecializePat (open_constr) (constr)" and "iSpecializePat_go", last call
failed.
Tactic failure: iSpecialize: cannot instantiate (⌜φ⌝ → P -∗ False)%I with P.
1 subgoal
PROP : sbi
P, Q : PROP
n, m, k : nat
============================
--------------------------------------∗
▷^(S n + S m) emp
"test_specialize_nested_intuitionistic"
: string
1 subgoal
PROP : sbi
φ : Prop
P, P2, Q, R1, R2 : PROP
H : φ
============================
"HP" : P
"HQ" : P -∗ Q
--------------------------------------□
"HR" : R2
--------------------------------------∗
R2
"test_iSimpl_in"
: string
1 subgoal
PROP : sbi
x, y : nat
============================
"H" : ⌜S (S (S x)) = y⌝
--------------------------------------∗
⌜S (S (S x)) = y⌝
1 subgoal
PROP : sbi
x, y, z : nat
============================
"H1" : ⌜S (S (S x)) = y⌝
"H2" : ⌜S y = z⌝
--------------------------------------∗
⌜S (S (S x)) = y⌝
1 subgoal
PROP : sbi
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:
In nested Ltac calls to "iSimpl in (constr)",
"iEval (tactic3) in (constr)" and "<iris.proofmode.ltac_tactics.iEval_go>",
last call failed.
Tactic failure: iEval: %: unsupported selection pattern.
1 subgoal
PROP : sbi
P, Q : PROP
============================
--------------------------------------∗
▷ emp
1 subgoal
PROP : sbi
P, Q : PROP
============================
--------------------------------------∗
▷ emp
The command has indeed failed with message:
In nested Ltac calls to "iFrame (constr)",
"<iris.proofmode.ltac_tactics.iFrame_go>" and
"<iris.proofmode.ltac_tactics.iFrameHyp>", last call failed.
Tactic failure: iFrame: cannot frame Q.
1 subgoal
PROP : sbi
BiAffine0 : BiAffine PROP
P, Q : PROP
============================
_ : □ P
_ : Q
--------------------------------------∗
□ P
1 subgoal
PROP : sbi
x : nat
l : list nat
P : PROP
============================
"HP" : P
_ : [∗ list] y ∈ l, <affine> ⌜y = y⌝
_ : [∗ list] y ∈ (x :: l), <affine> ⌜y = y⌝
--------------------------------------∗
P
1 subgoal
PROP : sbi
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
1 subgoal
PROP : sbi
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)
1 subgoal
PROP : sbi
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)
1 subgoal
PROP : sbi
Φ : 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
1 subgoal
PROP : sbi
============================
"H" : □ True
--------------------------------------∗
True
1 subgoal
PROP : sbi
============================
"H" : emp
--------------------------------------□
□ emp
"test_reducing_after_iApply_late_evar"
: string
1 subgoal
PROP : sbi
============================
"H" : emp
--------------------------------------□
□ emp
1 subgoal
PROP : sbi
mP : option PROP
Q, R : PROP
============================
"HPQ" : mP -∗? Q
"HQR" : Q -∗ R
--------------------------------------∗
R
1 subgoal
PROP : sbi
mP : option PROP
Q, R : PROP
============================
--------------------------------------∗
default emp mP
"elim_mod_accessor"
: string
1 subgoal
PROP : sbi
BiFUpd0 : BiFUpd PROP
X : Type
E1, E2 : coPset.coPset
α : X → PROP
β : X → PROP
γ : X → option PROP
============================
"Hacc" : ∃ x : X, α x ∗ (β x ={E2,E1}=∗ default emp (γ x))
--------------------------------------∗
|={E2,E1}=> True
1 subgoal
PROP : sbi
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
--------------------------------------∗
1 subgoal
PROP : sbi
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
--------------------------------------∗
1 subgoal
PROP : sbi
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 }}
--------------------------------------∗
1 subgoal
PROP : sbi
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 }}
--------------------------------------∗
1 subgoal
PROP : sbi
BiFUpd0 : BiFUpd PROP
PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
============================
--------------------------------------∗
PPPPPPPPPPPPPPPPP
→ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
1 subgoal
PROP : sbi
BiFUpd0 : BiFUpd PROP
PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
============================
--------------------------------------∗
PPPPPPPPPPPPPPPPP
→ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
→ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
1 subgoal
PROP : sbi
BiFUpd0 : BiFUpd PROP
PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
============================
--------------------------------------∗
PPPPPPPPPPPPPPPPP
-∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
1 subgoal
PROP : sbi
BiFUpd0 : BiFUpd PROP
PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
============================
--------------------------------------∗
PPPPPPPPPPPPPPPPP
-∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
-∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
1 subgoal
PROP : sbi
BiFUpd0 : BiFUpd PROP
E : coPset.coPset
PPPPPPPPPPPPPPPPP, QQQQQQQQQQQQQQQQQQ : PROP
============================
--------------------------------------∗
PPPPPPPPPPPPPPPPP
={E}=∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ ∗ QQQQQQQQQQQQQQQQQQ
1 subgoal
PROP : sbi
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:
Ltac call to "iStopProof" failed.
Tactic failure: iStopProof: proofmode not started.
"iAlways_spatial_non_empty"
: string
The command has indeed failed with message:
In nested Ltac calls to "iAlways", "iModIntro" and
"iModIntro (uconstr)", last call failed.
Tactic failure: iModIntro: spatial context is non-empty.
"iDestruct_bad_name"
: string
The command has indeed failed with message:
In nested Ltac calls to "iDestruct (open_constr) as (constr)",
"iDestructCore (open_constr) as (constr) (tactic3)" and
"iDestructCore (open_constr) as (constr) (tactic3)", last call failed.
Tactic failure: iDestruct: "HQ" not found.
"iIntros_dup_name"
: string
The command has indeed failed with message:
In nested Ltac calls to "iIntros (constr)", "iIntros_go" and
"iIntro (constr)", last call failed.
Tactic failure: iIntro: "HP" not fresh.
The command has indeed failed with message:
In nested Ltac calls to "iIntros ( (intropattern) )",
"iIntro ( (intropattern) )" and "intros x", last call failed.
x is already used.
"iSplitL_one_of_many"
The command has indeed failed with message:
Ltac call to "iSplitL (constr)" failed.
Tactic failure: iSplitL: hypotheses ["HPx"] not found.
The command has indeed failed with message:
Ltac call to "iSplitL (constr)" failed.
Tactic failure: iSplitL: hypotheses ["HPx"] not found.
"iSplitR_one_of_many"
: string
The command has indeed failed with message:
Ltac call to "iSplitR (constr)" failed.
Tactic failure: iSplitR: hypotheses ["HPx"] not found.
The command has indeed failed with message:
Ltac call to "iSplitR (constr)" failed.
Tactic failure: iSplitR: hypotheses ["HPx"] not found.
"iSplitL_non_splittable"
: string
The command has indeed failed with message:
Ltac call to "iSplitL (constr)" failed.
Tactic failure: iSplitL: P not a separating conjunction.
"iSplitR_non_splittable"
: string
The command has indeed failed with message:
Ltac call to "iSplitR (constr)" failed.
Tactic failure: iSplitR: P not a separating conjunction.
"iCombine_fail"
: string
The command has indeed failed with message:
Ltac call to "iCombine (constr) as (constr)" failed.
Tactic failure: iCombine: hypotheses ["HP3"] not found.
"iSpecialize_bad_name1_fail"
: string
The command has indeed failed with message:
In nested Ltac calls to "iSpecialize (open_constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) with (open_constr) (open_constr) as (constr)",
"iSpecializePat (open_constr) (constr)" and "iSpecializePat_go", last call
failed.
Tactic failure: iSpecialize: "H" not found.
"iSpecialize_bad_name2_fail"
: string
The command has indeed failed with message:
In nested Ltac calls to "iSpecialize (open_constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) with (open_constr) (open_constr) as (constr)",
"iSpecializePat (open_constr) (constr)" and "iSpecializePat_go", last call
failed.
Tactic failure: iSpecialize: "H" not found.
The command has indeed failed with message:
Ltac call to "iExact (constr)" failed.
Tactic failure: iExact: "HQ" not found.
520
521
522
523
524
525
526
527
528
529
530
531
532
533
534
535
536
537
538
539
540
541
542
543
544
545
546
547
548
549
550
551
552
553
554
555
556
557
558
559
The command has indeed failed with message:
Ltac call to "iExact (constr)" failed.
Tactic failure: iExact: "HQ" : Q does not match goal.
The command has indeed failed with message:
Ltac call to "iExact (constr)" failed.
Tactic failure: iExact: "HP"
not absorbing and the remaining hypotheses not affine.
"iClear_fail"
: string
The command has indeed failed with message:
In nested Ltac calls to "iClear (constr)", "iElaborateSelPat" and
"<iris.proofmode.ltac_tactics.iElaborateSelPat_go>", last call failed.
Tactic failure: iElaborateSelPat: "HP" not found.
The command has indeed failed with message:
In nested Ltac calls to "iClear (constr)",
"<iris.proofmode.ltac_tactics.iClear_go>" and
"<iris.proofmode.ltac_tactics.iClearHyp>", last call failed.
Tactic failure: iClear: "HP" : P not affine and the goal not absorbing.
"iSpecializeArgs_fail"
: string
The command has indeed failed with message:
In nested Ltac calls to "iSpecialize (open_constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) with (open_constr) (open_constr) as (constr)",
"iSpecializeArgs (constr) (open_constr)",
"<iris.proofmode.ltac_tactics.iSpecializeArgs_go>" and
"notypeclasses refine (uconstr)", last call failed.
In environment
PROP : sbi
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:
In nested Ltac calls to "iStartProof" and "iStartProof", last call failed.
Tactic failure: iStartProof: not a BI assertion.
"iPoseProof_fail"
: string
The command has indeed failed with message:

Ralf Jung
committed
In nested Ltac calls to "iPoseProof (open_constr) as (constr)",
"iPoseProofCore (open_constr) as (constr) (tactic3)",
"iPoseProofCoreLem (constr) as (tactic3)" and

Ralf Jung
committed
"<iris.proofmode.ltac_tactics.iIntoEmpValid>", last call failed.
Tactic failure: iPoseProof: not a BI assertion.
The command has indeed failed with message:

Ralf Jung
committed
In nested Ltac calls to "iPoseProof (open_constr) as (constr)",
"iPoseProofCore (open_constr) as (constr) (tactic3)",
"iPoseProofCoreLem (constr) as (tactic3)", "tac" (bound to
fun Htmp => spec_tac Htmp; [ .. | tac Htmp ]), "tac" (bound to
fun Htmp => spec_tac Htmp; [ .. | tac Htmp ]), "tac" (bound to
fun H => iDestructHyp H as pat), "iDestructHyp (constr) as (constr)",

Ralf Jung
committed
"<iris.proofmode.ltac_tactics.iDestructHypFindPat>",
"<iris.proofmode.ltac_tactics.iDestructHypGo>" and
"iRename (constr) into (constr)", last call failed.
Tactic failure: iRename: "H" not fresh.
"iPoseProof_not_found_fail"
: string
The command has indeed failed with message:
In nested Ltac calls to "iPoseProof (open_constr) as (constr)",
"iPoseProofCore (open_constr) as (constr) (tactic3)" and
"iPoseProofCoreHyp (constr) as (constr)", last call failed.
Tactic failure: iPoseProof: "Hx" not found.
"iPoseProof_not_found_fail2"
: string
The command has indeed failed with message:
In nested Ltac calls to "iPoseProof (open_constr) as (constr)",
"iPoseProofCore (open_constr) as (constr) (tactic3)",
"iPoseProofCoreLem (constr) as (tactic3)", "tac" (bound to
fun Htmp => spec_tac Htmp; [ .. | tac Htmp ]), "tac" (bound to
fun Htmp => spec_tac Htmp; [ .. | tac Htmp ]), "spec_tac" (bound to
fun Htmp =>
lazymatch lem with
| {| itrm := _; itrm_vars := ?xs; itrm_hyps := ?pat |} => iSpecializeCore
{| itrm := Htmp; itrm_vars := xs; itrm_hyps := pat |} as p
| _ => idtac
end), "iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) as (constr)",
"iSpecializeCore (open_constr) with (open_constr) (open_constr) as (constr)",
"iSpecializePat (open_constr) (constr)" and "iSpecializePat_go", last call
failed.
Tactic failure: iSpecialize: hypotheses ["HQ"] not found.
"iPoseProofCoreHyp_not_found_fail"
: string
The command has indeed failed with message:
Ltac call to "iPoseProofCoreHyp (constr) as (constr)" failed.
Tactic failure: iPoseProof: "Hx" not found.
"iPoseProofCoreHyp_not_fresh_fail"
: string
The command has indeed failed with message:
Ltac call to "iPoseProofCoreHyp (constr) as (constr)" failed.
Tactic failure: iPoseProof: "H1" not fresh.
"iRevert_fail"
: string
The command has indeed failed with message:
In nested Ltac calls to "iRevert (constr)", "iElaborateSelPat" and
"<iris.proofmode.ltac_tactics.iElaborateSelPat_go>", last call failed.
Tactic failure: iElaborateSelPat: "H" not found.
"iDestruct_fail"
: string
The command has indeed failed with message:
In nested Ltac calls to "iDestruct (open_constr) as (constr)",
"iDestructCore (open_constr) as (constr) (tactic3)",
"iDestructCore (open_constr) as (constr) (tactic3)",
"iDestructCore (open_constr) as (constr) (tactic3)",

Ralf Jung
committed
"tac" (bound to fun H => iDestructHyp H as pat),
"iDestructHyp (constr) as (constr)",
"<iris.proofmode.ltac_tactics.iDestructHypFindPat>" and
"<iris.proofmode.ltac_tactics.iDestructHypFindPat>", last call failed.
Tactic failure: iDestruct: "{HP}"
should contain exactly one proper introduction pattern.
The command has indeed failed with message:
In nested Ltac calls to "iDestruct (open_constr) as (constr)",
"iDestructCore (open_constr) as (constr) (tactic3)",
"iDestructCore (open_constr) as (constr) (tactic3)",
"iDestructCore (open_constr) as (constr) (tactic3)",

Ralf Jung
committed
"tac" (bound to fun H => iDestructHyp H as pat),
"iDestructHyp (constr) as (constr)",
"<iris.proofmode.ltac_tactics.iDestructHypFindPat>" and
"<iris.proofmode.ltac_tactics.iDestructHypGo>", last call failed.
Tactic failure: iDestruct: (IList [[IClear (sel_patterns.SelIdent "HP")]])
invalid.
"iOrDestruct_fail"
: string
The command has indeed failed with message:
Ltac call to "iOrDestruct (constr) as (constr) (constr)" failed.
Tactic failure: iOrDestruct: "H'" or "H2" not fresh.
The command has indeed failed with message:
Ltac call to "iOrDestruct (constr) as (constr) (constr)" failed.
Tactic failure: iOrDestruct: "H1" or "H'" not fresh.
"iApply_fail"
: string
The command has indeed failed with message:
In nested Ltac calls to "iApply (open_constr)",
"iPoseProofCore (open_constr) as (constr) (tactic3)",

Ralf Jung
committed
"tac" (bound to fun H => iApplyHyp H) and "iApplyHyp (constr)", last call
failed.
Tactic failure: iApply: cannot apply P.
"iApply_fail_not_affine_1"
: string
The command has indeed failed with message:
In nested Ltac calls to "iApply (open_constr)",
"iPoseProofCore (open_constr) as (constr) (tactic3)",
"tac" (bound to fun H => iApplyHyp H) and "iApplyHyp (constr)", last call
failed.
Tactic failure: iApply: Q
not absorbing and the remaining hypotheses not affine.
"iApply_fail_not_affine_2"
: string
The command has indeed failed with message:
In nested Ltac calls to "iApply (open_constr)",
"iPoseProofCore (open_constr) as (constr) (tactic3)",
"tac" (bound to fun H => iApplyHyp H) and "iApplyHyp (constr)", last call
failed.
Tactic failure: iApply: Q
not absorbing and the remaining hypotheses not affine.
676
677
678
679
680
681
682
683
684
685
686
687
688
689
690
691
692
693
694
695
696
697
698
699
700
701
702
703
704
705
706
707
708
709
710
711
712
713
"iRevert_wrong_var"
: string
The command has indeed failed with message:
In nested Ltac calls to "iRevert ( (ident) )" and
"iForallRevert (ident)", last call failed.
Tactic failure: iRevert: k1 not in scope.
The command has indeed failed with message:
In nested Ltac calls to "iLöb as (constr) forall ( (ident) )",
"iLöbRevert (constr) with (tactic3)",
"iRevertIntros (constr) with (tactic3)", "iRevertIntros_go",
"tac" (bound to iRevertIntros "∗" with tac), "tac" (bound to iRevertIntros
"∗" with tac), "iRevertIntros (constr) with (tactic3)",
"iRevertIntros_go", "tac" (bound to iRevertIntros ( x1 ) "" with iLöbCore as
IH), "tac" (bound to iRevertIntros ( x1 ) "" with iLöbCore as IH),
"iRevertIntros ( (ident) ) (constr) with (tactic3)",
"iRevertIntros (constr) with (tactic3)", "iRevertIntros_go",
"tac" (bound to iRevert ( x1 ); tac; iIntros ( x1 )),
"tac" (bound to iRevert ( x1 ); tac; iIntros ( x1 )),
"iRevert ( (ident) )" and "iForallRevert (ident)", last call failed.
Tactic failure: iRevert: k1 not in scope.
"iRevert_dup_var"
: string
The command has indeed failed with message:
In nested Ltac calls to "iRevert ( (ident) (ident) )",
"iRevert ( (ident) )" and "iForallRevert (ident)", last call failed.
Tactic failure: iRevert: k not in scope.
"iRevert_dep_var_coq"
: string
The command has indeed failed with message:
In nested Ltac calls to "iRevert ( (ident) )", "iForallRevert (ident)" and
"revert (ne_var_list)", last call failed.
k is used in hypothesis Hk.
"iRevert_dep_var"
: string
The command has indeed failed with message:
In nested Ltac calls to "iRevert ( (ident) )" and
"iForallRevert (ident)", last call failed.
Tactic failure: iRevert: k is used in hypothesis "Hk".
"iLöb_no_sbi"
: string
The command has indeed failed with message:
In nested Ltac calls to "iLöb as (constr)",
"iLöbRevert (constr) with (tactic3)",
"iRevertIntros (constr) with (tactic3)", "iRevertIntros_go",
"tac" (bound to iRevertIntros "∗" with tac), "tac" (bound to iRevertIntros
"∗" with tac), "iRevertIntros (constr) with (tactic3)",
"iRevertIntros_go", "tac" (bound to iLöbCore as IH),
"tac" (bound to iLöbCore as IH) and "iLöbCore as (constr)", last call failed.
Tactic failure: iLöb: not a step-indexed BI entailment.