Skip to content
Snippets Groups Projects
Commit 6de30407 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Fix potential stack overflow related to `Pretty N`.

parent 0a38be45
No related branches found
No related tags found
No related merge requests found
From stdpp Require Import pretty. From stdpp Require Import pretty.
From Coq Require Import Ascii.
Section N. Section N.
Local Open Scope N_scope. Local Open Scope N_scope.
...@@ -17,6 +18,17 @@ Section N. ...@@ -17,6 +18,17 @@ Section N.
Proof. reflexivity. Qed. Proof. reflexivity. Qed.
End N. End N.
(** Minimized version of:
https://coq.zulipchat.com/#narrow/stream/237977-Coq-users/topic/Stack.20overflow.20in.20Qed.2E
Fixed by making the [wp_guard] in [pretty_N_go] proportional to the
size of the input so that it blocks in case the input is an open term. *)
Lemma test_no_stack_overflow p n :
get n (pretty (N.pos p)) Some "_"%char
get (S n) ("-" +:+ pretty (N.pos p)) Some "_"%char.
Proof. intros Hlem. apply Hlem. Qed.
Section nat. Section nat.
Local Open Scope nat_scope. Local Open Scope nat_scope.
......
...@@ -23,7 +23,7 @@ Fixpoint pretty_N_go_help (x : N) (acc : Acc (<)%N x) (s : string) : string := ...@@ -23,7 +23,7 @@ Fixpoint pretty_N_go_help (x : N) (acc : Acc (<)%N x) (s : string) : string :=
| right _ => s | right _ => s
end. end.
Definition pretty_N_go (x : N) : string string := Definition pretty_N_go (x : N) : string string :=
pretty_N_go_help x (wf_guard 32 N.lt_wf_0 x). pretty_N_go_help x (wf_guard (S (N.size_nat x)) N.lt_wf_0 x).
Global Instance pretty_N : Pretty N := λ x, Global Instance pretty_N : Pretty N := λ x,
if decide (x = 0)%N then "0" else pretty_N_go x "". if decide (x = 0)%N then "0" else pretty_N_go x "".
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment