Skip to content
GitLab
Projects
Groups
Snippets
/
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Tej Chajed
stdpp
Commits
2c0cc639
Verified
Commit
2c0cc639
authored
Nov 06, 2020
by
Tej Chajed
Browse files
Pretty-print 0 as "0" for N, Z, and nat
Formerly printed as an empty string.
parent
1b559061
Changes
2
Hide whitespace changes
Inline
Side-by-side
CHANGELOG.md
View file @
2c0cc639
...
...
@@ -36,6 +36,8 @@ Coq 8.8 and 8.9 are no longer supported.
iterated addition.
+
Rename and restate many lemmas so as to be consistent with the conventions
for Coq's number types
`nat`
,
`N`
, and
`Z`
.
-
Fix a bug where
`pretty 0`
was defined as
`""`
, the empty string. It now
returns
`"0"`
for
`N`
,
`Z`
, and
`nat`
.
The following
`sed`
script should perform most of the renaming
(on macOS, replace
`sed`
by
`gsed`
, installed via e.g.
`brew install gnu-sed`
):
...
...
theories/pretty.v
View file @
2c0cc639
...
...
@@ -35,8 +35,8 @@ Proof.
unfold
pretty_N_go_help
at
1
;
fold
pretty_N_go_help
.
by
destruct
(
decide
(
0
<
x
)%
N
)
;
auto
using
pretty_N_go_help_irrel
.
Qed
.
Instance
pretty_N
:
Pretty
N
:
=
λ
x
,
pretty_N_go
x
""
%
string
.
Lemma
pretty_N_unfold
x
:
pretty
x
=
pretty_N_go
x
""
.
Instance
pretty_N
:
Pretty
N
:
=
λ
x
,
pretty_N_go
x
"
0
"
%
string
.
Lemma
pretty_N_unfold
x
:
pretty
x
=
pretty_N_go
x
"
0
"
.
Proof
.
done
.
Qed
.
Instance
pretty_N_inj
:
Inj
(=@{
N
})
(=)
pretty
.
Proof
.
...
...
@@ -46,7 +46,7 @@ Proof.
cut
(
∀
x
y
s
s'
,
pretty_N_go
x
s
=
pretty_N_go
y
s'
→
String
.
length
s
=
String
.
length
s'
→
x
=
y
∧
s
=
s'
).
{
intros
help
x
y
Hp
.
eapply
(
help
x
y
""
""
)
;
[
by
rewrite
<-!
pretty_N_unfold
|
done
].
}
eapply
(
help
x
y
"
0
"
"
0
"
)
;
[
by
rewrite
<-!
pretty_N_unfold
|
done
].
}
assert
(
∀
x
s
,
¬
String
.
length
(
pretty_N_go
x
s
)
<
String
.
length
s
)
as
help
.
{
setoid_rewrite
<-
Nat
.
le_ngt
.
intros
x
;
induction
(
N
.
lt_wf_0
x
)
as
[
x
_
IH
]
;
intros
s
.
...
...
@@ -66,9 +66,8 @@ Proof.
Qed
.
Instance
pretty_Z
:
Pretty
Z
:
=
λ
x
,
match
x
with
|
Z0
=>
""
|
Zpos
x
=>
pretty
(
Npos
x
)
|
Zneg
x
=>
"-"
+
:
+
pretty
(
Npos
x
)
|
Z0
=>
"
0
"
|
Zpos
x
=>
pretty
(
Npos
x
)
|
Zneg
x
=>
"-"
+
:
+
pretty
(
Npos
x
)
end
%
string
.
Instance
pretty_nat
:
Pretty
nat
:
=
λ
x
,
pretty
(
N
.
of_nat
x
).
Instance
pretty_nat_inj
:
Inj
(=@{
nat
})
(=)
pretty
.
Proof
.
apply
_
.
Qed
.
Write
Preview
Supports
Markdown
0%
Try again
or
attach a new file
.
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Cancel
Please
register
or
sign in
to comment