Skip to content

Commit 91f0945

Browse files
committed
fix #184
1 parent 7ee8e71 commit 91f0945

File tree

4 files changed

+41
-4
lines changed

4 files changed

+41
-4
lines changed

src/mmpars.c

Lines changed: 11 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -7,6 +7,7 @@
77
#include <string.h>
88
#include <stdlib.h>
99
#include <ctype.h>
10+
#include <limits.h>
1011
#include "mmvstr.h"
1112
#include "mmdata.h"
1213
#include "mminou.h"
@@ -2730,6 +2731,12 @@ char parseProof(long statemNum)
27302731
return returnFlag;
27312732
} // parseProof()
27322733

2734+
long saturatingMulAdd(long x, long y, long z) {
2735+
if (x > LONG_MAX / y) return LONG_MAX;
2736+
x *= y;
2737+
if (x > LONG_MAX - z) return LONG_MAX;
2738+
return x + z;
2739+
}
27332740
// Parse proof in compressed format.
27342741
// Parse proof of one statement in source file. Uses wrkProof structure.
27352742
// Returns 0 if OK; returns 1 if proof is incomplete (is empty or has '?'
@@ -3099,8 +3106,8 @@ char parseCompressedProof(long statemNum)
30993106
// * for each character c:
31003107
// * if c in ['U'..'Y']: n := n * 5 + (c - 'U' + 1)
31013108
// * if c in ['A'..'T']: n := n * 20 + (c - 'A' + 1)
3102-
labelMapIndex = labelMapIndex * lettersLen +
3103-
chrWeight[(long)(fbPtr[0])];
3109+
labelMapIndex = saturatingMulAdd(
3110+
labelMapIndex, lettersLen, chrWeight[(long)(fbPtr[0])]);
31043111
if (labelMapIndex >= g_WrkProof.compressedPfNumLabels) {
31053112
if (!g_WrkProof.errorCount) {
31063113
sourceError(labelStart, tokLength, statemNum, cat(
@@ -3184,8 +3191,8 @@ char parseCompressedProof(long statemNum)
31843191
labelMapIndex = chrWeight[(long)(fbPtr[0])] + 1;
31853192
labelStart = fbPtr; // Save label start for error msg
31863193
} else {
3187-
labelMapIndex = labelMapIndex * digitsLen +
3188-
chrWeight[(long)(fbPtr[0])] + 1;
3194+
labelMapIndex = saturatingMulAdd(
3195+
labelMapIndex, digitsLen, chrWeight[(long)(fbPtr[0])] + 1);
31893196
if (bggyAlgo) labelMapIndex--; // Adjust for buggy algorithm
31903197
}
31913198
break;

tests/issue184.expected

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,17 @@
1+
MM> READ "issue184.mm"
2+
Reading source file "issue184.mm"... 224 bytes
3+
224 bytes were read into the source buffer.
4+
The source has 6 statements; 0 are $a and 2 are $p.
5+
No errors were found. However, proofs were not checked. Type VERIFY PROOF *
6+
if you want to check them.
7+
MM> Continuous scrolling is now in effect.
8+
MM> 0 10% 20% 30% 40% 50% 60% 70% 80% 90% 100%
9+
..................................................
10+
?Error on line 12 of file "issue184.mm" at statement 6, label "evil", type
11+
"$p":
12+
( ww ) ABBBBBBVYVUXUUXYWYVVUUVUXWYVWYVYWG $.
13+
^^^^^^^^^^^^^^^^^^^^^^^^^^^
14+
This compressed label reference is outside the range of the label list. The
15+
compressed label value is 9223372036854775808 but the largest label defined is
16+
1.
17+

tests/issue184.in

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1 @@
1+
verify proof *

tests/issue184.mm

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,12 @@
1+
$c wff $.
2+
$c |- $.
3+
$v ph $.
4+
wph $f wff ph $.
5+
6+
$( A benign ` wff ` inference. $)
7+
ww $p wff ph $=
8+
( ) A $.
9+
10+
$( An evil ` |- ph ` proof. $)
11+
evil $p |- ph $=
12+
( ww ) ABBBBBBVYVUXUUXYWYVVUUVUXWYVWYVYWG $.

0 commit comments

Comments
 (0)