diff --git a/tests/heap_lang.v b/tests/heap_lang.v index 047e82d53b331b9653d8a11a0fdd2a54079b49b8..2d3f5047040fe1b1ba3651d108297dbcf74d6d77 100644 --- a/tests/heap_lang.v +++ b/tests/heap_lang.v @@ -115,6 +115,11 @@ Section tests. P -∗ (∀ Q Φ, Q -∗ WP e {{ Φ }}) -∗ WP e {{ _, True }}. Proof. iIntros "HP HW". wp_apply "HW". iExact "HP". Qed. +End tests. + +Section printing_tests. + Context `{heapG Σ}. + Lemma wp_print_long_expr (fun1 fun2 fun3 : expr) : True -∗ WP let: "val1" := fun1 #() in let: "val2" := fun2 "val1" in @@ -124,7 +129,7 @@ Section tests. iIntros "_". Show. Abort. -End tests. +End printing_tests. Lemma heap_e_adequate σ : adequate NotStuck heap_e σ (= #2). Proof. eapply (heap_adequacy heapΣ)=> ?. by apply heap_e_spec. Qed. diff --git a/tests/proofmode.ref b/tests/proofmode.ref index 28c05666197c1895e01afaaa8459f2948700110e..ae1266ae45527f2469cf0f655d9c5774711008e2 100644 --- a/tests/proofmode.ref +++ b/tests/proofmode.ref @@ -22,6 +22,19 @@ 1 subgoal PROP : sbi + P, Q : PROP + Persistent0 : Persistent P + Persistent1 : Persistent Q + ============================ + _ : P + _ : Q + --------------------------------------â–¡ + <affine> (P ∗ Q) + +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 @@ -32,6 +45,7 @@ 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 ============================ _ : 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 @@ -42,6 +56,7 @@ 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 : 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 | @@ -52,6 +67,7 @@ 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 : 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 | @@ -59,15 +75,3 @@ --------------------------------------∗ True -1 subgoal - - PROP : sbi - P, Q : PROP - Persistent0 : Persistent P - Persistent1 : Persistent Q - ============================ - _ : P - _ : Q - --------------------------------------â–¡ - <affine> (P ∗ Q) - diff --git a/tests/proofmode.v b/tests/proofmode.v index 63a7dd9f73279eed6998a2278eb98b9ff87b37e2..d5b83d55da0aaf40b1d55b01e7af502feea0c639 100644 --- a/tests/proofmode.v +++ b/tests/proofmode.v @@ -37,41 +37,6 @@ Lemma demo_3 P1 P2 P3 : P1 ∗ P2 ∗ P3 -∗ P1 ∗ â–· (P2 ∗ ∃ x, (P3 ∧ ⌜x = 0âŒ) ∨ P3). Proof. iIntros "($ & $ & $)". iNext. by iExists 0. Qed. -(* Test line breaking of long assumptions. *) -Section linebreaks. -Lemma print_long_line (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 - -∗ True. -Proof. - iIntros "HP". Show. -Abort. -Lemma print_long_line_anon (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 - -∗ True. -Proof. - iIntros "?". Show. -Abort. - -(* This is specifically crafted such that not having the `hv` in - the proofmode notation breaks the output. *) -Local Notation "'TESTNOTATION' '{{' P '|' Q '}' '}'" := (P ∧ Q)%I - (format "'TESTNOTATION' '{{' P '|' '/' Q '}' '}'") : bi_scope. -Lemma print_long_line (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 }} - -∗ True. -Proof. - iIntros "HP". Show. -Abort. -Lemma print_long_line_anon (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 }} - -∗ True. -Proof. - iIntros "?". Show. -Abort. -End linebreaks. - Definition foo (P : PROP) := (P -∗ P)%I. Definition bar : PROP := (∀ P, foo P)%I. @@ -513,3 +478,45 @@ Proof. lazymatch goal with |- coq_tactics.envs_entails _ (â–¡ P) => done end. Qed. End tests. + +(** Test specifically if certain things print correctly. *) +Section printing_tests. +Context {PROP : sbi} `{!BiFUpd PROP}. +Implicit Types P Q R : PROP. + +(* Test line breaking of long assumptions. *) +Section linebreaks. +Lemma print_long_line (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 + -∗ True. +Proof. + iIntros "HP". Show. +Abort. +Lemma print_long_line_anon (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 + -∗ True. +Proof. + iIntros "?". Show. +Abort. + +(* This is specifically crafted such that not having the `hv` in + the proofmode notation breaks the output. *) +Local Notation "'TESTNOTATION' '{{' P '|' Q '}' '}'" := (P ∧ Q)%I + (format "'TESTNOTATION' '{{' P '|' '/' Q '}' '}'") : bi_scope. +Lemma print_long_line (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 }} + -∗ True. +Proof. + iIntros "HP". Show. +Abort. +Lemma print_long_line_anon (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 }} + -∗ True. +Proof. + iIntros "?". Show. +Abort. +End linebreaks. + +End printing_tests.