mirror of
https://github.com/github/codeql.git
synced 2025-12-17 01:03:14 +01:00
Merge pull request #938 from dave-bartolomeo/dave/AliasedSSA
C++: Better tracking of SSA memory accesses
This commit is contained in:
@@ -68,6 +68,10 @@
|
||||
"cpp/ql/src/semmle/code/cpp/ir/implementation/unaliased_ssa/internal/SSAConstruction.qll",
|
||||
"cpp/ql/src/semmle/code/cpp/ir/implementation/aliased_ssa/internal/SSAConstruction.qll"
|
||||
],
|
||||
"C++ SSA PrintSSA": [
|
||||
"cpp/ql/src/semmle/code/cpp/ir/implementation/unaliased_ssa/internal/PrintSSA.qll",
|
||||
"cpp/ql/src/semmle/code/cpp/ir/implementation/aliased_ssa/internal/PrintSSA.qll"
|
||||
],
|
||||
"C++ IR ValueNumber": [
|
||||
"cpp/ql/src/semmle/code/cpp/ir/implementation/raw/gvn/ValueNumbering.qll",
|
||||
"cpp/ql/src/semmle/code/cpp/ir/implementation/unaliased_ssa/gvn/ValueNumbering.qll",
|
||||
|
||||
@@ -16,6 +16,13 @@ class Enum extends UserType, IntegralOrEnumType {
|
||||
*/
|
||||
override string explain() { result = "enum " + this.getName() }
|
||||
|
||||
override int getSize() {
|
||||
// Workaround for extractor bug CPP-348: No size information for enums.
|
||||
// If the extractor didn't provide a size, assume four bytes.
|
||||
result = UserType.super.getSize() or
|
||||
not exists(UserType.super.getSize()) and result = 4
|
||||
}
|
||||
|
||||
/** See `Type.isDeeplyConst` and `Type.isDeeplyConstBelow`. Internal. */
|
||||
override predicate isDeeplyConstBelow() { any() } // No subparts
|
||||
|
||||
|
||||
@@ -6,6 +6,7 @@ private newtype TMemoryAccessKind =
|
||||
TBufferMemoryAccess() or
|
||||
TBufferMayMemoryAccess() or
|
||||
TEscapedMemoryAccess() or
|
||||
TEscapedMayMemoryAccess() or
|
||||
TPhiMemoryAccess() or
|
||||
TUnmodeledMemoryAccess() or
|
||||
TChiTotalMemoryAccess() or
|
||||
@@ -16,7 +17,17 @@ private newtype TMemoryAccessKind =
|
||||
* memory result.
|
||||
*/
|
||||
class MemoryAccessKind extends TMemoryAccessKind {
|
||||
abstract string toString();
|
||||
string toString() {
|
||||
none()
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if the operand or result accesses memory pointed to by the `AddressOperand` on the
|
||||
* same instruction.
|
||||
*/
|
||||
predicate usesAddressOperand() {
|
||||
none()
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
@@ -27,6 +38,10 @@ class IndirectMemoryAccess extends MemoryAccessKind, TIndirectMemoryAccess {
|
||||
override string toString() {
|
||||
result = "indirect"
|
||||
}
|
||||
|
||||
override final predicate usesAddressOperand() {
|
||||
any()
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
@@ -37,6 +52,10 @@ class IndirectMayMemoryAccess extends MemoryAccessKind, TIndirectMayMemoryAccess
|
||||
override string toString() {
|
||||
result = "indirect(may)"
|
||||
}
|
||||
|
||||
override final predicate usesAddressOperand() {
|
||||
any()
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
@@ -48,6 +67,10 @@ class BufferMemoryAccess extends MemoryAccessKind, TBufferMemoryAccess {
|
||||
override string toString() {
|
||||
result = "buffer"
|
||||
}
|
||||
|
||||
override final predicate usesAddressOperand() {
|
||||
any()
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
@@ -59,6 +82,10 @@ class BufferMayMemoryAccess extends MemoryAccessKind, TBufferMayMemoryAccess {
|
||||
override string toString() {
|
||||
result = "buffer(may)"
|
||||
}
|
||||
|
||||
override final predicate usesAddressOperand() {
|
||||
any()
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
@@ -70,6 +97,15 @@ class EscapedMemoryAccess extends MemoryAccessKind, TEscapedMemoryAccess {
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
* The operand or result may access all memory whose address has escaped.
|
||||
*/
|
||||
class EscapedMayMemoryAccess extends MemoryAccessKind, TEscapedMayMemoryAccess {
|
||||
override string toString() {
|
||||
result = "escaped(may)"
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
* The operand is a Phi operand, which accesses the same memory as its
|
||||
* definition.
|
||||
|
||||
@@ -498,6 +498,16 @@ class Instruction extends Construction::TInstruction {
|
||||
none()
|
||||
}
|
||||
|
||||
/**
|
||||
* Returns the operand that holds the memory address to which the instruction stores its
|
||||
* result, if any. For example, in `m3 = Store r1, r2`, the result of `getResultAddressOperand()`
|
||||
* is `r1`.
|
||||
*/
|
||||
final AddressOperand getResultAddressOperand() {
|
||||
getResultMemoryAccess().usesAddressOperand() and
|
||||
result.getUseInstruction() = this
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if the result of this instruction is precisely modeled in SSA. Always
|
||||
* holds for a register result. For a memory result, a modeled result is
|
||||
@@ -1340,7 +1350,7 @@ class CallSideEffectInstruction extends SideEffectInstruction {
|
||||
}
|
||||
|
||||
override final MemoryAccessKind getResultMemoryAccess() {
|
||||
result instanceof EscapedMemoryAccess
|
||||
result instanceof EscapedMayMemoryAccess
|
||||
}
|
||||
}
|
||||
|
||||
|
||||
@@ -128,7 +128,7 @@ class MemoryOperand extends Operand {
|
||||
* is `r1`.
|
||||
*/
|
||||
final AddressOperand getAddressOperand() {
|
||||
getMemoryAccess() instanceof IndirectMemoryAccess and
|
||||
getMemoryAccess().usesAddressOperand() and
|
||||
result.getUseInstruction() = getUseInstruction()
|
||||
}
|
||||
}
|
||||
@@ -351,10 +351,10 @@ class SideEffectOperand extends TypedOperand {
|
||||
|
||||
override MemoryAccessKind getMemoryAccess() {
|
||||
useInstr instanceof CallSideEffectInstruction and
|
||||
result instanceof EscapedMemoryAccess
|
||||
result instanceof EscapedMayMemoryAccess
|
||||
or
|
||||
useInstr instanceof CallReadSideEffectInstruction and
|
||||
result instanceof EscapedMemoryAccess
|
||||
result instanceof EscapedMayMemoryAccess
|
||||
or
|
||||
useInstr instanceof IndirectReadSideEffectInstruction and
|
||||
result instanceof IndirectMemoryAccess
|
||||
|
||||
@@ -1,11 +1,10 @@
|
||||
import cpp
|
||||
import AliasAnalysis
|
||||
private import semmle.code.cpp.ir.implementation.unaliased_ssa.IR
|
||||
private import semmle.code.cpp.ir.internal.OperandTag
|
||||
private import semmle.code.cpp.ir.internal.Overlap
|
||||
|
||||
import semmle.code.cpp.ir.internal.Overlap
|
||||
private import semmle.code.cpp.ir.implementation.unaliased_ssa.IR
|
||||
private import semmle.code.cpp.ir.internal.IntegerConstant as Ints
|
||||
private import semmle.code.cpp.ir.internal.IntegerInterval as Interval
|
||||
private import semmle.code.cpp.ir.internal.OperandTag
|
||||
|
||||
private class IntValue = Ints::IntValue;
|
||||
|
||||
@@ -37,6 +36,9 @@ class VirtualVariable extends TVirtualVariable {
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
* A virtual variable representing a single non-escaped `IRVariable`.
|
||||
*/
|
||||
class VirtualIRVariable extends VirtualVariable, TVirtualIRVariable {
|
||||
IRVariable var;
|
||||
|
||||
@@ -52,7 +54,6 @@ class VirtualIRVariable extends VirtualVariable, TVirtualIRVariable {
|
||||
result = var
|
||||
}
|
||||
|
||||
// REVIEW: This should just be on MemoryAccess
|
||||
override final Type getType() {
|
||||
result = var.getType()
|
||||
}
|
||||
@@ -62,6 +63,10 @@ class VirtualIRVariable extends VirtualVariable, TVirtualIRVariable {
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
* A virtual variable representing all escaped memory accessible by the function,
|
||||
* including escaped local variables.
|
||||
*/
|
||||
class UnknownVirtualVariable extends VirtualVariable, TUnknownVirtualVariable {
|
||||
FunctionIR f;
|
||||
|
||||
@@ -86,24 +91,35 @@ class UnknownVirtualVariable extends VirtualVariable, TUnknownVirtualVariable {
|
||||
}
|
||||
}
|
||||
|
||||
private predicate hasResultMemoryAccess(Instruction instr, IRVariable var, IntValue startBitOffset,
|
||||
IntValue endBitOffset) {
|
||||
resultPointsTo(instr.getResultAddressOperand().getDefinitionInstruction(), var, startBitOffset) and
|
||||
if exists(instr.getResultSize()) then
|
||||
endBitOffset = Ints::add(startBitOffset, Ints::mul(instr.getResultSize(), 8))
|
||||
else
|
||||
endBitOffset = Ints::unknown()
|
||||
}
|
||||
|
||||
private predicate hasOperandMemoryAccess(MemoryOperand operand, IRVariable var, IntValue startBitOffset,
|
||||
IntValue endBitOffset) {
|
||||
resultPointsTo(operand.getAddressOperand().getDefinitionInstruction(), var, startBitOffset) and
|
||||
if exists(operand.getSize()) then
|
||||
endBitOffset = Ints::add(startBitOffset, Ints::mul(operand.getSize(), 8))
|
||||
else
|
||||
endBitOffset = Ints::unknown()
|
||||
}
|
||||
|
||||
private newtype TMemoryAccess =
|
||||
TVariableMemoryAccess(IRVariable var, IntValue offset, IntValue size) {
|
||||
exists(Instruction instr |
|
||||
exists(MemoryAccessKind mak | instr.getResultMemoryAccess() = mak and not mak instanceof PhiMemoryAccess) and
|
||||
resultPointsTo(instr.getAnOperand().(AddressOperand).getDefinitionInstruction(), var, offset) and
|
||||
if exists(instr.getResultSize())
|
||||
then instr.getResultSize() = size
|
||||
else size = Ints::unknown()
|
||||
)
|
||||
TVariableMemoryAccess(IRVariable var, IntValue startBitOffset, IntValue endBitOffset) {
|
||||
hasResultMemoryAccess(_, var, startBitOffset, endBitOffset) or
|
||||
hasOperandMemoryAccess(_, var, startBitOffset, endBitOffset)
|
||||
}
|
||||
or
|
||||
TUnknownMemoryAccess(UnknownVirtualVariable uvv) or
|
||||
TTotalUnknownMemoryAccess(UnknownVirtualVariable uvv)
|
||||
|
||||
private VariableMemoryAccess getVariableMemoryAccess(IRVariable var, IntValue offset, IntValue size) {
|
||||
result.getVariable() = var and
|
||||
result.getOffset() = offset and
|
||||
result.getSize() = size
|
||||
private VariableMemoryAccess getVariableMemoryAccess(IRVariable var, IntValue startBitOffset, IntValue endBitOffset) {
|
||||
result = TVariableMemoryAccess(var, startBitOffset, endBitOffset)
|
||||
}
|
||||
|
||||
class MemoryAccess extends TMemoryAccess {
|
||||
@@ -120,17 +136,27 @@ class MemoryAccess extends TMemoryAccess {
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
* An access to memory within a single known `IRVariable`. The variable may be either an unescaped variable
|
||||
* (with its own `VirtualIRVariable`) or an escaped variable (assiged to `UnknownVirtualVariable`).
|
||||
*/
|
||||
class VariableMemoryAccess extends TVariableMemoryAccess, MemoryAccess {
|
||||
IRVariable var;
|
||||
IntValue offset;
|
||||
IntValue size;
|
||||
IntValue startBitOffset;
|
||||
IntValue endBitOffset;
|
||||
|
||||
VariableMemoryAccess() {
|
||||
this = TVariableMemoryAccess(var, offset, size)
|
||||
this = TVariableMemoryAccess(var, startBitOffset, endBitOffset)
|
||||
}
|
||||
|
||||
override final string toString() {
|
||||
result = var.toString() + "[" + offset.toString() + ".." + (offset + size - 1).toString() + "]"
|
||||
exists(string partialString |
|
||||
result = var.toString() + Interval::getIntervalString(startBitOffset, endBitOffset) + partialString and
|
||||
if isPartialMemoryAccess() then
|
||||
partialString = " (partial)"
|
||||
else
|
||||
partialString = ""
|
||||
)
|
||||
}
|
||||
|
||||
final override VirtualVariable getVirtualVariable() {
|
||||
@@ -138,12 +164,12 @@ class VariableMemoryAccess extends TVariableMemoryAccess, MemoryAccess {
|
||||
not exists(getVirtualVariable(var)) and result = getUnknownVirtualVariable(var.getEnclosingFunctionIR())
|
||||
}
|
||||
|
||||
IntValue getOffset() {
|
||||
result = offset
|
||||
IntValue getStartBitOffset() {
|
||||
result = startBitOffset
|
||||
}
|
||||
|
||||
IntValue getSize() {
|
||||
result = size
|
||||
IntValue getEndBitOffset() {
|
||||
result = endBitOffset
|
||||
}
|
||||
|
||||
final IRVariable getVariable() {
|
||||
@@ -152,12 +178,15 @@ class VariableMemoryAccess extends TVariableMemoryAccess, MemoryAccess {
|
||||
|
||||
final override predicate isPartialMemoryAccess() {
|
||||
not exists(getVirtualVariable(var)) or
|
||||
getOffset() != 0
|
||||
getStartBitOffset() != 0
|
||||
or
|
||||
getSize() != var.getType().getSize()
|
||||
not Ints::isEQ(getEndBitOffset(), Ints::add(getStartBitOffset(), Ints::mul(var.getType().getSize(), 8)))
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
* An access to memory that is not known to be confined to a specific `IRVariable`.
|
||||
*/
|
||||
class UnknownMemoryAccess extends TUnknownMemoryAccess, MemoryAccess {
|
||||
UnknownVirtualVariable vvar;
|
||||
|
||||
@@ -176,12 +205,11 @@ class UnknownMemoryAccess extends TUnknownMemoryAccess, MemoryAccess {
|
||||
final override predicate isPartialMemoryAccess() {
|
||||
any()
|
||||
}
|
||||
|
||||
Type getType() {
|
||||
result instanceof UnknownType
|
||||
}
|
||||
}
|
||||
|
||||
/**
|
||||
* An access to all aliased memory.
|
||||
*/
|
||||
class TotalUnknownMemoryAccess extends TTotalUnknownMemoryAccess, MemoryAccess {
|
||||
UnknownVirtualVariable vvar;
|
||||
|
||||
@@ -196,84 +224,91 @@ class TotalUnknownMemoryAccess extends TTotalUnknownMemoryAccess, MemoryAccess {
|
||||
final override VirtualVariable getVirtualVariable() {
|
||||
result = vvar
|
||||
}
|
||||
|
||||
Type getType() {
|
||||
result instanceof UnknownType
|
||||
}
|
||||
}
|
||||
|
||||
Overlap getOverlap(MemoryAccess def, MemoryAccess use) {
|
||||
def instanceof VariableMemoryAccess and
|
||||
def = use and
|
||||
result instanceof MustExactlyOverlap
|
||||
or
|
||||
exists(VariableMemoryAccess defVMA, VariableMemoryAccess useVMA, int defOffset, int defEnd,
|
||||
int useOffset, int useEnd |
|
||||
defVMA = def and
|
||||
useVMA = use and
|
||||
defVMA.getVirtualVariable() = useVMA.getVirtualVariable() and
|
||||
defVMA != useVMA and
|
||||
defOffset = Ints::getValue(defVMA.getOffset()) and
|
||||
defEnd = Ints::getValue(Ints::add(defVMA.getOffset(), defVMA.getSize())) and
|
||||
useOffset = Ints::getValue(useVMA.getOffset()) and
|
||||
useEnd = Ints::getValue(Ints::add(useVMA.getOffset(), useVMA.getSize()))
|
||||
|
|
||||
defOffset <= useOffset and
|
||||
defEnd >= useEnd and
|
||||
result instanceof MustTotallyOverlap
|
||||
or
|
||||
defOffset > useOffset and
|
||||
defOffset < useEnd and
|
||||
result instanceof MayPartiallyOverlap
|
||||
or
|
||||
defOffset = useOffset and
|
||||
defEnd < useEnd and
|
||||
result instanceof MayPartiallyOverlap
|
||||
)
|
||||
or
|
||||
exists(UnknownVirtualVariable uvv |
|
||||
def = TUnknownMemoryAccess(uvv) and
|
||||
uvv = use.getVirtualVariable() and
|
||||
result instanceof MayPartiallyOverlap
|
||||
)
|
||||
or
|
||||
exists(UnknownVirtualVariable uvv |
|
||||
def = TTotalUnknownMemoryAccess(uvv) and
|
||||
uvv = use.getVirtualVariable() and
|
||||
result instanceof MustTotallyOverlap
|
||||
def.getVirtualVariable() = use.getVirtualVariable() and
|
||||
(
|
||||
// A TotalUnknownMemoryAccess must totally overlap any access to the same virtual variable.
|
||||
def instanceof TotalUnknownMemoryAccess and result instanceof MustTotallyOverlap or
|
||||
// An UnknownMemoryAccess may partially overlap any access to the same virtual variable.
|
||||
def instanceof UnknownMemoryAccess and result instanceof MayPartiallyOverlap or
|
||||
exists(VariableMemoryAccess defVariableAccess |
|
||||
defVariableAccess = def and
|
||||
(
|
||||
(
|
||||
// A VariableMemoryAccess may partially overlap an unknown access to the same virtual variable.
|
||||
((use instanceof UnknownMemoryAccess) or (use instanceof TotalUnknownMemoryAccess)) and
|
||||
result instanceof MayPartiallyOverlap
|
||||
) or
|
||||
// A VariableMemoryAccess overlaps another access to the same variable based on the relationship
|
||||
// of the two offset intervals.
|
||||
exists(VariableMemoryAccess useVariableAccess, IntValue defStartOffset, IntValue defEndOffset,
|
||||
IntValue useStartOffset, IntValue useEndOffset |
|
||||
useVariableAccess = use and
|
||||
defStartOffset = defVariableAccess.getStartBitOffset() and
|
||||
defEndOffset = defVariableAccess.getEndBitOffset() and
|
||||
useStartOffset = useVariableAccess.getStartBitOffset() and
|
||||
useEndOffset = useVariableAccess.getEndBitOffset() and
|
||||
result = Interval::getOverlap(defStartOffset, defEndOffset, useStartOffset, useEndOffset)
|
||||
)
|
||||
)
|
||||
)
|
||||
)
|
||||
}
|
||||
|
||||
MemoryAccess getResultMemoryAccess(Instruction instr) {
|
||||
exists(instr.getResultMemoryAccess()) and
|
||||
if exists(IRVariable var, IntValue i |
|
||||
resultPointsTo(instr.getAnOperand().(AddressOperand).getDefinitionInstruction(), var, i)
|
||||
)
|
||||
then exists(IRVariable var, IntValue i |
|
||||
resultPointsTo(instr.getAnOperand().(AddressOperand).getDefinitionInstruction(), var, i) and
|
||||
result = getVariableMemoryAccess(var, i, instr.getResultSize())
|
||||
)
|
||||
else (
|
||||
result = TUnknownMemoryAccess(TUnknownVirtualVariable(instr.getEnclosingFunctionIR())) and
|
||||
not instr instanceof UnmodeledDefinitionInstruction and
|
||||
not instr instanceof AliasedDefinitionInstruction
|
||||
or
|
||||
result = TTotalUnknownMemoryAccess(TUnknownVirtualVariable(instr.getEnclosingFunctionIR())) and
|
||||
instr instanceof AliasedDefinitionInstruction
|
||||
exists(MemoryAccessKind kind |
|
||||
kind = instr.getResultMemoryAccess() and
|
||||
(
|
||||
(
|
||||
kind.usesAddressOperand() and
|
||||
if hasResultMemoryAccess(instr, _, _, _) then (
|
||||
exists(IRVariable var, IntValue startBitOffset, IntValue endBitOffset |
|
||||
hasResultMemoryAccess(instr, var, startBitOffset, endBitOffset) and
|
||||
result = getVariableMemoryAccess(var, startBitOffset, endBitOffset)
|
||||
)
|
||||
)
|
||||
else (
|
||||
result = TUnknownMemoryAccess(TUnknownVirtualVariable(instr.getEnclosingFunctionIR()))
|
||||
)
|
||||
) or
|
||||
(
|
||||
kind instanceof EscapedMemoryAccess and
|
||||
result = TTotalUnknownMemoryAccess(TUnknownVirtualVariable(instr.getEnclosingFunctionIR()))
|
||||
) or
|
||||
(
|
||||
kind instanceof EscapedMayMemoryAccess and
|
||||
result = TUnknownMemoryAccess(TUnknownVirtualVariable(instr.getEnclosingFunctionIR()))
|
||||
)
|
||||
)
|
||||
)
|
||||
}
|
||||
|
||||
MemoryAccess getOperandMemoryAccess(MemoryOperand operand) {
|
||||
if exists(IRVariable var, IntValue i |
|
||||
resultPointsTo(operand.getAddressOperand().getDefinitionInstruction(), var, i)
|
||||
)
|
||||
then exists(IRVariable var, IntValue i, int size |
|
||||
resultPointsTo(operand.getAddressOperand().getDefinitionInstruction(), var, i) and
|
||||
result = getVariableMemoryAccess(var, i, size) and
|
||||
size = operand.getDefinitionInstruction().getResultSize()
|
||||
)
|
||||
else (
|
||||
result = TUnknownMemoryAccess(TUnknownVirtualVariable(operand.getUseInstruction().getEnclosingFunctionIR())) and
|
||||
not operand.getUseInstruction() instanceof UnmodeledUseInstruction
|
||||
exists(MemoryAccessKind kind |
|
||||
kind = operand.getMemoryAccess() and
|
||||
(
|
||||
(
|
||||
kind.usesAddressOperand() and
|
||||
if hasOperandMemoryAccess(operand, _, _, _) then (
|
||||
exists(IRVariable var, IntValue startBitOffset, IntValue endBitOffset |
|
||||
hasOperandMemoryAccess(operand, var, startBitOffset, endBitOffset) and
|
||||
result = getVariableMemoryAccess(var, startBitOffset, endBitOffset)
|
||||
)
|
||||
)
|
||||
else (
|
||||
result = TUnknownMemoryAccess(TUnknownVirtualVariable(operand.getEnclosingFunctionIR()))
|
||||
)
|
||||
) or
|
||||
(
|
||||
kind instanceof EscapedMemoryAccess and
|
||||
result = TTotalUnknownMemoryAccess(TUnknownVirtualVariable(operand.getEnclosingFunctionIR()))
|
||||
) or
|
||||
(
|
||||
kind instanceof EscapedMayMemoryAccess and
|
||||
result = TUnknownMemoryAccess(TUnknownVirtualVariable(operand.getEnclosingFunctionIR()))
|
||||
)
|
||||
)
|
||||
)
|
||||
}
|
||||
|
||||
@@ -1,12 +0,0 @@
|
||||
private import semmle.code.cpp.ir.implementation.unaliased_ssa.IR
|
||||
private import AliasedSSA
|
||||
|
||||
/**
|
||||
* Property provide that dumps the memory access of each result. Useful for debugging SSA
|
||||
* construction.
|
||||
*/
|
||||
class PropertyProvider extends IRPropertyProvider {
|
||||
override string getInstructionProperty(Instruction instruction, string key) {
|
||||
key = "ResultMemoryAccess" and result = getResultMemoryAccess(instruction).toString()
|
||||
}
|
||||
}
|
||||
@@ -0,0 +1,20 @@
|
||||
private import SSAConstructionInternal
|
||||
private import OldIR
|
||||
private import Alias
|
||||
|
||||
/**
|
||||
* Property provide that dumps the memory access of each result. Useful for debugging SSA
|
||||
* construction.
|
||||
*/
|
||||
class PropertyProvider extends IRPropertyProvider {
|
||||
override string getInstructionProperty(Instruction instruction, string key) {
|
||||
(
|
||||
key = "ResultMemoryAccess" and
|
||||
result = getResultMemoryAccess(instruction).toString()
|
||||
) or
|
||||
(
|
||||
key = "OperandMemoryAccess" and
|
||||
result = getOperandMemoryAccess(instruction.getAnOperand().(MemoryOperand)).toString()
|
||||
)
|
||||
}
|
||||
}
|
||||
@@ -498,6 +498,16 @@ class Instruction extends Construction::TInstruction {
|
||||
none()
|
||||
}
|
||||
|
||||
/**
|
||||
* Returns the operand that holds the memory address to which the instruction stores its
|
||||
* result, if any. For example, in `m3 = Store r1, r2`, the result of `getResultAddressOperand()`
|
||||
* is `r1`.
|
||||
*/
|
||||
final AddressOperand getResultAddressOperand() {
|
||||
getResultMemoryAccess().usesAddressOperand() and
|
||||
result.getUseInstruction() = this
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if the result of this instruction is precisely modeled in SSA. Always
|
||||
* holds for a register result. For a memory result, a modeled result is
|
||||
@@ -1340,7 +1350,7 @@ class CallSideEffectInstruction extends SideEffectInstruction {
|
||||
}
|
||||
|
||||
override final MemoryAccessKind getResultMemoryAccess() {
|
||||
result instanceof EscapedMemoryAccess
|
||||
result instanceof EscapedMayMemoryAccess
|
||||
}
|
||||
}
|
||||
|
||||
|
||||
@@ -128,7 +128,7 @@ class MemoryOperand extends Operand {
|
||||
* is `r1`.
|
||||
*/
|
||||
final AddressOperand getAddressOperand() {
|
||||
getMemoryAccess() instanceof IndirectMemoryAccess and
|
||||
getMemoryAccess().usesAddressOperand() and
|
||||
result.getUseInstruction() = getUseInstruction()
|
||||
}
|
||||
}
|
||||
@@ -351,10 +351,10 @@ class SideEffectOperand extends TypedOperand {
|
||||
|
||||
override MemoryAccessKind getMemoryAccess() {
|
||||
useInstr instanceof CallSideEffectInstruction and
|
||||
result instanceof EscapedMemoryAccess
|
||||
result instanceof EscapedMayMemoryAccess
|
||||
or
|
||||
useInstr instanceof CallReadSideEffectInstruction and
|
||||
result instanceof EscapedMemoryAccess
|
||||
result instanceof EscapedMayMemoryAccess
|
||||
or
|
||||
useInstr instanceof IndirectReadSideEffectInstruction and
|
||||
result instanceof IndirectMemoryAccess
|
||||
|
||||
@@ -498,6 +498,16 @@ class Instruction extends Construction::TInstruction {
|
||||
none()
|
||||
}
|
||||
|
||||
/**
|
||||
* Returns the operand that holds the memory address to which the instruction stores its
|
||||
* result, if any. For example, in `m3 = Store r1, r2`, the result of `getResultAddressOperand()`
|
||||
* is `r1`.
|
||||
*/
|
||||
final AddressOperand getResultAddressOperand() {
|
||||
getResultMemoryAccess().usesAddressOperand() and
|
||||
result.getUseInstruction() = this
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if the result of this instruction is precisely modeled in SSA. Always
|
||||
* holds for a register result. For a memory result, a modeled result is
|
||||
@@ -1340,7 +1350,7 @@ class CallSideEffectInstruction extends SideEffectInstruction {
|
||||
}
|
||||
|
||||
override final MemoryAccessKind getResultMemoryAccess() {
|
||||
result instanceof EscapedMemoryAccess
|
||||
result instanceof EscapedMayMemoryAccess
|
||||
}
|
||||
}
|
||||
|
||||
|
||||
@@ -128,7 +128,7 @@ class MemoryOperand extends Operand {
|
||||
* is `r1`.
|
||||
*/
|
||||
final AddressOperand getAddressOperand() {
|
||||
getMemoryAccess() instanceof IndirectMemoryAccess and
|
||||
getMemoryAccess().usesAddressOperand() and
|
||||
result.getUseInstruction() = getUseInstruction()
|
||||
}
|
||||
}
|
||||
@@ -351,10 +351,10 @@ class SideEffectOperand extends TypedOperand {
|
||||
|
||||
override MemoryAccessKind getMemoryAccess() {
|
||||
useInstr instanceof CallSideEffectInstruction and
|
||||
result instanceof EscapedMemoryAccess
|
||||
result instanceof EscapedMayMemoryAccess
|
||||
or
|
||||
useInstr instanceof CallReadSideEffectInstruction and
|
||||
result instanceof EscapedMemoryAccess
|
||||
result instanceof EscapedMayMemoryAccess
|
||||
or
|
||||
useInstr instanceof IndirectReadSideEffectInstruction and
|
||||
result instanceof IndirectMemoryAccess
|
||||
|
||||
@@ -0,0 +1,20 @@
|
||||
private import SSAConstructionInternal
|
||||
private import OldIR
|
||||
private import Alias
|
||||
|
||||
/**
|
||||
* Property provide that dumps the memory access of each result. Useful for debugging SSA
|
||||
* construction.
|
||||
*/
|
||||
class PropertyProvider extends IRPropertyProvider {
|
||||
override string getInstructionProperty(Instruction instruction, string key) {
|
||||
(
|
||||
key = "ResultMemoryAccess" and
|
||||
result = getResultMemoryAccess(instruction).toString()
|
||||
) or
|
||||
(
|
||||
key = "OperandMemoryAccess" and
|
||||
result = getOperandMemoryAccess(instruction.getAnOperand().(MemoryOperand)).toString()
|
||||
)
|
||||
}
|
||||
}
|
||||
@@ -1,12 +1,46 @@
|
||||
import AliasAnalysis
|
||||
import cpp
|
||||
private import semmle.code.cpp.ir.implementation.raw.IR
|
||||
private import semmle.code.cpp.ir.internal.IntegerConstant as Ints
|
||||
private import semmle.code.cpp.ir.internal.OperandTag
|
||||
private import semmle.code.cpp.ir.internal.Overlap
|
||||
|
||||
private class IntValue = Ints::IntValue;
|
||||
|
||||
private predicate hasResultMemoryAccess(Instruction instr, IRVariable var, Type type, IntValue bitOffset) {
|
||||
resultPointsTo(instr.getResultAddressOperand().getDefinitionInstruction(), var, bitOffset) and
|
||||
type = instr.getResultType()
|
||||
}
|
||||
|
||||
private predicate hasOperandMemoryAccess(MemoryOperand operand, IRVariable var, Type type, IntValue bitOffset) {
|
||||
resultPointsTo(operand.getAddressOperand().getDefinitionInstruction(), var, bitOffset) and
|
||||
type = operand.getType()
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if the specified variable should be modeled in SSA form. For unaliased SSA, we only model a variable if its
|
||||
* address never escapes and all reads and writes of that variable access the entire variable using the original type
|
||||
* of the variable.
|
||||
*/
|
||||
private predicate isVariableModeled(IRVariable var) {
|
||||
not variableAddressEscapes(var) and
|
||||
// There's no need to check for the right size. An `IRVariable` never has an `UnknownType`, so the test for
|
||||
// `type = var.getType()` is sufficient.
|
||||
forall(Instruction instr, Type type, IntValue bitOffset |
|
||||
hasResultMemoryAccess(instr, var, type, bitOffset) |
|
||||
bitOffset = 0 and
|
||||
type = var.getType()
|
||||
) and
|
||||
forall(MemoryOperand operand, Type type, IntValue bitOffset |
|
||||
hasOperandMemoryAccess(operand, var, type, bitOffset) |
|
||||
bitOffset = 0 and
|
||||
type = var.getType()
|
||||
)
|
||||
}
|
||||
|
||||
private newtype TVirtualVariable =
|
||||
MkVirtualVariable(IRVariable var) {
|
||||
not variableAddressEscapes(var)
|
||||
isVariableModeled(var)
|
||||
}
|
||||
|
||||
private VirtualVariable getVirtualVariable(IRVariable var) {
|
||||
@@ -28,7 +62,6 @@ class VirtualVariable extends TVirtualVariable {
|
||||
result = var
|
||||
}
|
||||
|
||||
// REVIEW: This should just be on MemoryAccess
|
||||
final Type getType() {
|
||||
result = var.getType()
|
||||
}
|
||||
@@ -74,16 +107,14 @@ Overlap getOverlap(MemoryAccess def, MemoryAccess use) {
|
||||
|
||||
MemoryAccess getResultMemoryAccess(Instruction instr) {
|
||||
exists(IRVariable var |
|
||||
instr.getResultMemoryAccess() instanceof IndirectMemoryAccess and
|
||||
resultPointsTo(instr.getAnOperand().(AddressOperand).getDefinitionInstruction(),
|
||||
var, 0) and
|
||||
hasResultMemoryAccess(instr, var, _, _) and
|
||||
result = getMemoryAccess(var)
|
||||
)
|
||||
}
|
||||
|
||||
MemoryAccess getOperandMemoryAccess(MemoryOperand operand) {
|
||||
exists(IRVariable var |
|
||||
resultPointsTo(operand.getAddressOperand().getDefinitionInstruction(), var, 0) and
|
||||
hasOperandMemoryAccess(operand, var, _, _) and
|
||||
result = getMemoryAccess(var)
|
||||
)
|
||||
}
|
||||
|
||||
@@ -31,6 +31,14 @@ predicate hasValue(IntValue n) {
|
||||
n != unknown()
|
||||
}
|
||||
|
||||
/**
|
||||
* Returns a string representation of `n`. If `n` does not have a known value, the result is "??".
|
||||
*/
|
||||
bindingset[n]
|
||||
string intValueToString(IntValue n) {
|
||||
if hasValue(n) then result = n.toString() else result = "??"
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if the value `f` is within the range of representable integers.
|
||||
*/
|
||||
@@ -199,3 +207,51 @@ bindingset[a]
|
||||
IntValue neg(IntValue a) {
|
||||
result = -a // -INT_MIN = INT_MIN, so this preserves unknown
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if `a` is equal to `b`. Does not hold if either `a` or `b` is unknown.
|
||||
*/
|
||||
bindingset[a, b]
|
||||
predicate isEQ(IntValue a, IntValue b) {
|
||||
hasValue(a) and hasValue(b) and a = b
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if `a` is not equal to `b`. Does not hold if either `a` or `b` is unknown.
|
||||
*/
|
||||
bindingset[a, b]
|
||||
predicate isNE(IntValue a, IntValue b) {
|
||||
hasValue(a) and hasValue(b) and a != b
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if `a` is less than `b`. Does not hold if either `a` or `b` is unknown.
|
||||
*/
|
||||
bindingset[a, b]
|
||||
predicate isLT(IntValue a, IntValue b) {
|
||||
hasValue(a) and hasValue(b) and a < b
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if `a` is less than or equal to `b`. Does not hold if either `a` or `b` is unknown.
|
||||
*/
|
||||
bindingset[a, b]
|
||||
predicate isLE(IntValue a, IntValue b) {
|
||||
hasValue(a) and hasValue(b) and a <= b
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if `a` is greater than `b`. Does not hold if either `a` or `b` is unknown.
|
||||
*/
|
||||
bindingset[a, b]
|
||||
predicate isGT(IntValue a, IntValue b) {
|
||||
hasValue(a) and hasValue(b) and a > b
|
||||
}
|
||||
|
||||
/**
|
||||
* Holds if `a` is greater than or equal to `b`. Does not hold if either `a` or `b` is unknown.
|
||||
*/
|
||||
bindingset[a, b]
|
||||
predicate isGE(IntValue a, IntValue b) {
|
||||
hasValue(a) and hasValue(b) and a >= b
|
||||
}
|
||||
|
||||
33
cpp/ql/src/semmle/code/cpp/ir/internal/IntegerInterval.qll
Normal file
33
cpp/ql/src/semmle/code/cpp/ir/internal/IntegerInterval.qll
Normal file
@@ -0,0 +1,33 @@
|
||||
/**
|
||||
* Support for integer intervals.
|
||||
* An interval is represented as by its inclusive lower bound, `start`, and its exclusive upper bound, `end`.
|
||||
* Either or both of `start` and `end` may have an unknown value.
|
||||
*/
|
||||
|
||||
import Overlap
|
||||
private import IntegerConstant
|
||||
|
||||
/**
|
||||
* Gets the overlap relationship between the definition interval [`defStart`, `defEnd`) and the use interval
|
||||
* [`useStart`, `useEnd`).
|
||||
*/
|
||||
bindingset[defStart, defEnd, useStart, useEnd]
|
||||
Overlap getOverlap(IntValue defStart, IntValue defEnd, IntValue useStart, IntValue useEnd) {
|
||||
if isEQ(defStart, useStart) and isEQ(defEnd, useEnd) then
|
||||
result instanceof MustExactlyOverlap
|
||||
else if isLE(defStart, useStart) and isGE(defEnd, useEnd) then
|
||||
result instanceof MustTotallyOverlap
|
||||
else if isLE(defEnd, useStart) or isGE(defStart, useEnd) then
|
||||
none()
|
||||
else
|
||||
result instanceof MayPartiallyOverlap
|
||||
}
|
||||
|
||||
/**
|
||||
* Gets a string representation of the interval [`start`, `end`).
|
||||
*/
|
||||
bindingset[start, end]
|
||||
string getIntervalString(IntValue start, IntValue end) {
|
||||
// We represent an interval has half-open, so print it as "[start..end)".
|
||||
result = "[" + intValueToString(start) + ".." + intValueToString(end) + ")"
|
||||
}
|
||||
@@ -1,6 +1,7 @@
|
||||
| test.cpp:66:30:66:36 | test.cpp:71:8:71:9 | AST only |
|
||||
| test.cpp:89:28:89:34 | test.cpp:92:8:92:14 | IR only |
|
||||
| test.cpp:100:13:100:18 | test.cpp:103:10:103:12 | AST only |
|
||||
| test.cpp:109:9:109:14 | test.cpp:110:10:110:12 | IR only |
|
||||
| test.cpp:120:9:120:20 | test.cpp:126:8:126:19 | AST only |
|
||||
| test.cpp:122:18:122:30 | test.cpp:132:22:132:23 | IR only |
|
||||
| test.cpp:122:18:122:30 | test.cpp:140:22:140:23 | IR only |
|
||||
|
||||
@@ -11,6 +11,7 @@
|
||||
| test.cpp:86:8:86:9 | Load: i1 | test.cpp:83:7:83:8 | Uninitialized: definition of u2 |
|
||||
| test.cpp:90:8:90:14 | Load: source1 | test.cpp:89:28:89:34 | InitializeParameter: source1 |
|
||||
| test.cpp:92:8:92:14 | Load: source1 | test.cpp:89:28:89:34 | InitializeParameter: source1 |
|
||||
| test.cpp:110:10:110:12 | Load: (reference dereference) | test.cpp:109:9:109:14 | Call: call to source |
|
||||
| test.cpp:132:22:132:23 | Load: m1 | test.cpp:122:18:122:30 | InitializeParameter: sourceStruct1 |
|
||||
| test.cpp:140:22:140:23 | Load: m1 | test.cpp:122:18:122:30 | InitializeParameter: sourceStruct1 |
|
||||
| test.cpp:188:8:188:8 | Load: y | test.cpp:186:27:186:32 | Call: call to source |
|
||||
|
||||
@@ -235,3 +235,160 @@ ssa.cpp:
|
||||
# 69| v3_10(void) = ConditionalBranch : r3_9
|
||||
#-----| False -> Block 2
|
||||
#-----| True -> Block 1
|
||||
|
||||
# 75| void MustExactlyOverlap(Point)
|
||||
# 75| Block 0
|
||||
# 75| v0_0(void) = EnterFunction :
|
||||
# 75| m0_1(unknown) = AliasedDefinition :
|
||||
# 75| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 75| r0_3(glval<Point>) = VariableAddress[a] :
|
||||
# 75| m0_4(Point) = InitializeParameter[a] : r0_3
|
||||
# 76| r0_5(glval<Point>) = VariableAddress[b] :
|
||||
# 76| r0_6(glval<Point>) = VariableAddress[a] :
|
||||
# 76| r0_7(Point) = Load : r0_6, m0_4
|
||||
# 76| m0_8(Point) = Store : r0_5, r0_7
|
||||
# 77| v0_9(void) = NoOp :
|
||||
# 75| v0_10(void) = ReturnVoid :
|
||||
# 75| v0_11(void) = UnmodeledUse : mu*
|
||||
# 75| v0_12(void) = ExitFunction :
|
||||
|
||||
# 79| void MustExactlyOverlapEscaped(Point)
|
||||
# 79| Block 0
|
||||
# 79| v0_0(void) = EnterFunction :
|
||||
# 79| m0_1(unknown) = AliasedDefinition :
|
||||
# 79| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 79| r0_3(glval<Point>) = VariableAddress[a] :
|
||||
# 79| m0_4(Point) = InitializeParameter[a] : r0_3
|
||||
# 79| m0_5(unknown) = Chi : m0_1, m0_4
|
||||
# 80| r0_6(glval<Point>) = VariableAddress[b] :
|
||||
# 80| r0_7(glval<Point>) = VariableAddress[a] :
|
||||
# 80| r0_8(Point) = Load : r0_7, m0_5
|
||||
# 80| m0_9(Point) = Store : r0_6, r0_8
|
||||
# 81| r0_10(glval<unknown>) = FunctionAddress[Escape] :
|
||||
# 81| r0_11(glval<Point>) = VariableAddress[a] :
|
||||
# 81| r0_12(void *) = Convert : r0_11
|
||||
# 81| v0_13(void) = Call : r0_10, r0_12
|
||||
# 81| m0_14(unknown) = ^CallSideEffect : m0_5
|
||||
# 81| m0_15(unknown) = Chi : m0_5, m0_14
|
||||
# 82| v0_16(void) = NoOp :
|
||||
# 79| v0_17(void) = ReturnVoid :
|
||||
# 79| v0_18(void) = UnmodeledUse : mu*
|
||||
# 79| v0_19(void) = ExitFunction :
|
||||
|
||||
# 84| void MustTotallyOverlap(Point)
|
||||
# 84| Block 0
|
||||
# 84| v0_0(void) = EnterFunction :
|
||||
# 84| m0_1(unknown) = AliasedDefinition :
|
||||
# 84| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 84| r0_3(glval<Point>) = VariableAddress[a] :
|
||||
# 84| m0_4(Point) = InitializeParameter[a] : r0_3
|
||||
# 85| r0_5(glval<int>) = VariableAddress[x] :
|
||||
# 85| r0_6(glval<Point>) = VariableAddress[a] :
|
||||
# 85| r0_7(glval<int>) = FieldAddress[x] : r0_6
|
||||
# 85| r0_8(int) = Load : r0_7, m0_4
|
||||
# 85| m0_9(int) = Store : r0_5, r0_8
|
||||
# 86| r0_10(glval<int>) = VariableAddress[y] :
|
||||
# 86| r0_11(glval<Point>) = VariableAddress[a] :
|
||||
# 86| r0_12(glval<int>) = FieldAddress[y] : r0_11
|
||||
# 86| r0_13(int) = Load : r0_12, m0_4
|
||||
# 86| m0_14(int) = Store : r0_10, r0_13
|
||||
# 87| v0_15(void) = NoOp :
|
||||
# 84| v0_16(void) = ReturnVoid :
|
||||
# 84| v0_17(void) = UnmodeledUse : mu*
|
||||
# 84| v0_18(void) = ExitFunction :
|
||||
|
||||
# 89| void MustTotallyOverlapEscaped(Point)
|
||||
# 89| Block 0
|
||||
# 89| v0_0(void) = EnterFunction :
|
||||
# 89| m0_1(unknown) = AliasedDefinition :
|
||||
# 89| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 89| r0_3(glval<Point>) = VariableAddress[a] :
|
||||
# 89| m0_4(Point) = InitializeParameter[a] : r0_3
|
||||
# 89| m0_5(unknown) = Chi : m0_1, m0_4
|
||||
# 90| r0_6(glval<int>) = VariableAddress[x] :
|
||||
# 90| r0_7(glval<Point>) = VariableAddress[a] :
|
||||
# 90| r0_8(glval<int>) = FieldAddress[x] : r0_7
|
||||
# 90| r0_9(int) = Load : r0_8, m0_5
|
||||
# 90| m0_10(int) = Store : r0_6, r0_9
|
||||
# 91| r0_11(glval<int>) = VariableAddress[y] :
|
||||
# 91| r0_12(glval<Point>) = VariableAddress[a] :
|
||||
# 91| r0_13(glval<int>) = FieldAddress[y] : r0_12
|
||||
# 91| r0_14(int) = Load : r0_13, m0_5
|
||||
# 91| m0_15(int) = Store : r0_11, r0_14
|
||||
# 92| r0_16(glval<unknown>) = FunctionAddress[Escape] :
|
||||
# 92| r0_17(glval<Point>) = VariableAddress[a] :
|
||||
# 92| r0_18(void *) = Convert : r0_17
|
||||
# 92| v0_19(void) = Call : r0_16, r0_18
|
||||
# 92| m0_20(unknown) = ^CallSideEffect : m0_5
|
||||
# 92| m0_21(unknown) = Chi : m0_5, m0_20
|
||||
# 93| v0_22(void) = NoOp :
|
||||
# 89| v0_23(void) = ReturnVoid :
|
||||
# 89| v0_24(void) = UnmodeledUse : mu*
|
||||
# 89| v0_25(void) = ExitFunction :
|
||||
|
||||
# 95| void MayPartiallyOverlap(int, int)
|
||||
# 95| Block 0
|
||||
# 95| v0_0(void) = EnterFunction :
|
||||
# 95| m0_1(unknown) = AliasedDefinition :
|
||||
# 95| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 95| r0_3(glval<int>) = VariableAddress[x] :
|
||||
# 95| m0_4(int) = InitializeParameter[x] : r0_3
|
||||
# 95| r0_5(glval<int>) = VariableAddress[y] :
|
||||
# 95| m0_6(int) = InitializeParameter[y] : r0_5
|
||||
# 96| r0_7(glval<Point>) = VariableAddress[a] :
|
||||
# 96| m0_8(Point) = Uninitialized[a] : r0_7
|
||||
# 96| r0_9(glval<int>) = FieldAddress[x] : r0_7
|
||||
# 96| r0_10(glval<int>) = VariableAddress[x] :
|
||||
# 96| r0_11(int) = Load : r0_10, m0_4
|
||||
# 96| m0_12(int) = Store : r0_9, r0_11
|
||||
# 96| m0_13(Point) = Chi : m0_8, m0_12
|
||||
# 96| r0_14(glval<int>) = FieldAddress[y] : r0_7
|
||||
# 96| r0_15(glval<int>) = VariableAddress[y] :
|
||||
# 96| r0_16(int) = Load : r0_15, m0_6
|
||||
# 96| m0_17(int) = Store : r0_14, r0_16
|
||||
# 96| m0_18(Point) = Chi : m0_13, m0_17
|
||||
# 97| r0_19(glval<Point>) = VariableAddress[b] :
|
||||
# 97| r0_20(glval<Point>) = VariableAddress[a] :
|
||||
# 97| r0_21(Point) = Load : r0_20, m0_18
|
||||
# 97| m0_22(Point) = Store : r0_19, r0_21
|
||||
# 98| v0_23(void) = NoOp :
|
||||
# 95| v0_24(void) = ReturnVoid :
|
||||
# 95| v0_25(void) = UnmodeledUse : mu*
|
||||
# 95| v0_26(void) = ExitFunction :
|
||||
|
||||
# 100| void MayPartiallyOverlapEscaped(int, int)
|
||||
# 100| Block 0
|
||||
# 100| v0_0(void) = EnterFunction :
|
||||
# 100| m0_1(unknown) = AliasedDefinition :
|
||||
# 100| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 100| r0_3(glval<int>) = VariableAddress[x] :
|
||||
# 100| m0_4(int) = InitializeParameter[x] : r0_3
|
||||
# 100| r0_5(glval<int>) = VariableAddress[y] :
|
||||
# 100| m0_6(int) = InitializeParameter[y] : r0_5
|
||||
# 101| r0_7(glval<Point>) = VariableAddress[a] :
|
||||
# 101| m0_8(Point) = Uninitialized[a] : r0_7
|
||||
# 101| m0_9(unknown) = Chi : m0_1, m0_8
|
||||
# 101| r0_10(glval<int>) = FieldAddress[x] : r0_7
|
||||
# 101| r0_11(glval<int>) = VariableAddress[x] :
|
||||
# 101| r0_12(int) = Load : r0_11, m0_4
|
||||
# 101| m0_13(int) = Store : r0_10, r0_12
|
||||
# 101| m0_14(unknown) = Chi : m0_9, m0_13
|
||||
# 101| r0_15(glval<int>) = FieldAddress[y] : r0_7
|
||||
# 101| r0_16(glval<int>) = VariableAddress[y] :
|
||||
# 101| r0_17(int) = Load : r0_16, m0_6
|
||||
# 101| m0_18(int) = Store : r0_15, r0_17
|
||||
# 101| m0_19(unknown) = Chi : m0_14, m0_18
|
||||
# 102| r0_20(glval<Point>) = VariableAddress[b] :
|
||||
# 102| r0_21(glval<Point>) = VariableAddress[a] :
|
||||
# 102| r0_22(Point) = Load : r0_21, m0_19
|
||||
# 102| m0_23(Point) = Store : r0_20, r0_22
|
||||
# 103| r0_24(glval<unknown>) = FunctionAddress[Escape] :
|
||||
# 103| r0_25(glval<Point>) = VariableAddress[a] :
|
||||
# 103| r0_26(void *) = Convert : r0_25
|
||||
# 103| v0_27(void) = Call : r0_24, r0_26
|
||||
# 103| m0_28(unknown) = ^CallSideEffect : m0_19
|
||||
# 103| m0_29(unknown) = Chi : m0_19, m0_28
|
||||
# 104| v0_30(void) = NoOp :
|
||||
# 100| v0_31(void) = ReturnVoid :
|
||||
# 100| v0_32(void) = UnmodeledUse : mu*
|
||||
# 100| v0_33(void) = ExitFunction :
|
||||
|
||||
@@ -69,3 +69,36 @@ void chiNodeAtEndOfLoop(int n, char* p) {
|
||||
while (n-- > 0)
|
||||
* p++ = 0;
|
||||
}
|
||||
|
||||
void Escape(void* p);
|
||||
|
||||
void MustExactlyOverlap(Point a) {
|
||||
Point b = a;
|
||||
}
|
||||
|
||||
void MustExactlyOverlapEscaped(Point a) {
|
||||
Point b = a;
|
||||
Escape(&a);
|
||||
}
|
||||
|
||||
void MustTotallyOverlap(Point a) {
|
||||
int x = a.x;
|
||||
int y = a.y;
|
||||
}
|
||||
|
||||
void MustTotallyOverlapEscaped(Point a) {
|
||||
int x = a.x;
|
||||
int y = a.y;
|
||||
Escape(&a);
|
||||
}
|
||||
|
||||
void MayPartiallyOverlap(int x, int y) {
|
||||
Point a = { x, y };
|
||||
Point b = a;
|
||||
}
|
||||
|
||||
void MayPartiallyOverlapEscaped(int x, int y) {
|
||||
Point a = { x, y };
|
||||
Point b = a;
|
||||
Escape(&a);
|
||||
}
|
||||
|
||||
@@ -236,3 +236,150 @@ ssa.cpp:
|
||||
# 69| v3_9(void) = ConditionalBranch : r3_8
|
||||
#-----| False -> Block 2
|
||||
#-----| True -> Block 1
|
||||
|
||||
# 75| void MustExactlyOverlap(Point)
|
||||
# 75| Block 0
|
||||
# 75| v0_0(void) = EnterFunction :
|
||||
# 75| mu0_1(unknown) = AliasedDefinition :
|
||||
# 75| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 75| r0_3(glval<Point>) = VariableAddress[a] :
|
||||
# 75| m0_4(Point) = InitializeParameter[a] : r0_3
|
||||
# 76| r0_5(glval<Point>) = VariableAddress[b] :
|
||||
# 76| r0_6(glval<Point>) = VariableAddress[a] :
|
||||
# 76| r0_7(Point) = Load : r0_6, m0_4
|
||||
# 76| m0_8(Point) = Store : r0_5, r0_7
|
||||
# 77| v0_9(void) = NoOp :
|
||||
# 75| v0_10(void) = ReturnVoid :
|
||||
# 75| v0_11(void) = UnmodeledUse : mu*
|
||||
# 75| v0_12(void) = ExitFunction :
|
||||
|
||||
# 79| void MustExactlyOverlapEscaped(Point)
|
||||
# 79| Block 0
|
||||
# 79| v0_0(void) = EnterFunction :
|
||||
# 79| mu0_1(unknown) = AliasedDefinition :
|
||||
# 79| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 79| r0_3(glval<Point>) = VariableAddress[a] :
|
||||
# 79| mu0_4(Point) = InitializeParameter[a] : r0_3
|
||||
# 80| r0_5(glval<Point>) = VariableAddress[b] :
|
||||
# 80| r0_6(glval<Point>) = VariableAddress[a] :
|
||||
# 80| r0_7(Point) = Load : r0_6, mu0_2
|
||||
# 80| m0_8(Point) = Store : r0_5, r0_7
|
||||
# 81| r0_9(glval<unknown>) = FunctionAddress[Escape] :
|
||||
# 81| r0_10(glval<Point>) = VariableAddress[a] :
|
||||
# 81| r0_11(void *) = Convert : r0_10
|
||||
# 81| v0_12(void) = Call : r0_9, r0_11
|
||||
# 81| mu0_13(unknown) = ^CallSideEffect : mu0_2
|
||||
# 82| v0_14(void) = NoOp :
|
||||
# 79| v0_15(void) = ReturnVoid :
|
||||
# 79| v0_16(void) = UnmodeledUse : mu*
|
||||
# 79| v0_17(void) = ExitFunction :
|
||||
|
||||
# 84| void MustTotallyOverlap(Point)
|
||||
# 84| Block 0
|
||||
# 84| v0_0(void) = EnterFunction :
|
||||
# 84| mu0_1(unknown) = AliasedDefinition :
|
||||
# 84| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 84| r0_3(glval<Point>) = VariableAddress[a] :
|
||||
# 84| mu0_4(Point) = InitializeParameter[a] : r0_3
|
||||
# 85| r0_5(glval<int>) = VariableAddress[x] :
|
||||
# 85| r0_6(glval<Point>) = VariableAddress[a] :
|
||||
# 85| r0_7(glval<int>) = FieldAddress[x] : r0_6
|
||||
# 85| r0_8(int) = Load : r0_7, mu0_2
|
||||
# 85| m0_9(int) = Store : r0_5, r0_8
|
||||
# 86| r0_10(glval<int>) = VariableAddress[y] :
|
||||
# 86| r0_11(glval<Point>) = VariableAddress[a] :
|
||||
# 86| r0_12(glval<int>) = FieldAddress[y] : r0_11
|
||||
# 86| r0_13(int) = Load : r0_12, mu0_2
|
||||
# 86| m0_14(int) = Store : r0_10, r0_13
|
||||
# 87| v0_15(void) = NoOp :
|
||||
# 84| v0_16(void) = ReturnVoid :
|
||||
# 84| v0_17(void) = UnmodeledUse : mu*
|
||||
# 84| v0_18(void) = ExitFunction :
|
||||
|
||||
# 89| void MustTotallyOverlapEscaped(Point)
|
||||
# 89| Block 0
|
||||
# 89| v0_0(void) = EnterFunction :
|
||||
# 89| mu0_1(unknown) = AliasedDefinition :
|
||||
# 89| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 89| r0_3(glval<Point>) = VariableAddress[a] :
|
||||
# 89| mu0_4(Point) = InitializeParameter[a] : r0_3
|
||||
# 90| r0_5(glval<int>) = VariableAddress[x] :
|
||||
# 90| r0_6(glval<Point>) = VariableAddress[a] :
|
||||
# 90| r0_7(glval<int>) = FieldAddress[x] : r0_6
|
||||
# 90| r0_8(int) = Load : r0_7, mu0_2
|
||||
# 90| m0_9(int) = Store : r0_5, r0_8
|
||||
# 91| r0_10(glval<int>) = VariableAddress[y] :
|
||||
# 91| r0_11(glval<Point>) = VariableAddress[a] :
|
||||
# 91| r0_12(glval<int>) = FieldAddress[y] : r0_11
|
||||
# 91| r0_13(int) = Load : r0_12, mu0_2
|
||||
# 91| m0_14(int) = Store : r0_10, r0_13
|
||||
# 92| r0_15(glval<unknown>) = FunctionAddress[Escape] :
|
||||
# 92| r0_16(glval<Point>) = VariableAddress[a] :
|
||||
# 92| r0_17(void *) = Convert : r0_16
|
||||
# 92| v0_18(void) = Call : r0_15, r0_17
|
||||
# 92| mu0_19(unknown) = ^CallSideEffect : mu0_2
|
||||
# 93| v0_20(void) = NoOp :
|
||||
# 89| v0_21(void) = ReturnVoid :
|
||||
# 89| v0_22(void) = UnmodeledUse : mu*
|
||||
# 89| v0_23(void) = ExitFunction :
|
||||
|
||||
# 95| void MayPartiallyOverlap(int, int)
|
||||
# 95| Block 0
|
||||
# 95| v0_0(void) = EnterFunction :
|
||||
# 95| mu0_1(unknown) = AliasedDefinition :
|
||||
# 95| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 95| r0_3(glval<int>) = VariableAddress[x] :
|
||||
# 95| m0_4(int) = InitializeParameter[x] : r0_3
|
||||
# 95| r0_5(glval<int>) = VariableAddress[y] :
|
||||
# 95| m0_6(int) = InitializeParameter[y] : r0_5
|
||||
# 96| r0_7(glval<Point>) = VariableAddress[a] :
|
||||
# 96| mu0_8(Point) = Uninitialized[a] : r0_7
|
||||
# 96| r0_9(glval<int>) = FieldAddress[x] : r0_7
|
||||
# 96| r0_10(glval<int>) = VariableAddress[x] :
|
||||
# 96| r0_11(int) = Load : r0_10, m0_4
|
||||
# 96| mu0_12(int) = Store : r0_9, r0_11
|
||||
# 96| r0_13(glval<int>) = FieldAddress[y] : r0_7
|
||||
# 96| r0_14(glval<int>) = VariableAddress[y] :
|
||||
# 96| r0_15(int) = Load : r0_14, m0_6
|
||||
# 96| mu0_16(int) = Store : r0_13, r0_15
|
||||
# 97| r0_17(glval<Point>) = VariableAddress[b] :
|
||||
# 97| r0_18(glval<Point>) = VariableAddress[a] :
|
||||
# 97| r0_19(Point) = Load : r0_18, mu0_2
|
||||
# 97| m0_20(Point) = Store : r0_17, r0_19
|
||||
# 98| v0_21(void) = NoOp :
|
||||
# 95| v0_22(void) = ReturnVoid :
|
||||
# 95| v0_23(void) = UnmodeledUse : mu*
|
||||
# 95| v0_24(void) = ExitFunction :
|
||||
|
||||
# 100| void MayPartiallyOverlapEscaped(int, int)
|
||||
# 100| Block 0
|
||||
# 100| v0_0(void) = EnterFunction :
|
||||
# 100| mu0_1(unknown) = AliasedDefinition :
|
||||
# 100| mu0_2(unknown) = UnmodeledDefinition :
|
||||
# 100| r0_3(glval<int>) = VariableAddress[x] :
|
||||
# 100| m0_4(int) = InitializeParameter[x] : r0_3
|
||||
# 100| r0_5(glval<int>) = VariableAddress[y] :
|
||||
# 100| m0_6(int) = InitializeParameter[y] : r0_5
|
||||
# 101| r0_7(glval<Point>) = VariableAddress[a] :
|
||||
# 101| mu0_8(Point) = Uninitialized[a] : r0_7
|
||||
# 101| r0_9(glval<int>) = FieldAddress[x] : r0_7
|
||||
# 101| r0_10(glval<int>) = VariableAddress[x] :
|
||||
# 101| r0_11(int) = Load : r0_10, m0_4
|
||||
# 101| mu0_12(int) = Store : r0_9, r0_11
|
||||
# 101| r0_13(glval<int>) = FieldAddress[y] : r0_7
|
||||
# 101| r0_14(glval<int>) = VariableAddress[y] :
|
||||
# 101| r0_15(int) = Load : r0_14, m0_6
|
||||
# 101| mu0_16(int) = Store : r0_13, r0_15
|
||||
# 102| r0_17(glval<Point>) = VariableAddress[b] :
|
||||
# 102| r0_18(glval<Point>) = VariableAddress[a] :
|
||||
# 102| r0_19(Point) = Load : r0_18, mu0_2
|
||||
# 102| m0_20(Point) = Store : r0_17, r0_19
|
||||
# 103| r0_21(glval<unknown>) = FunctionAddress[Escape] :
|
||||
# 103| r0_22(glval<Point>) = VariableAddress[a] :
|
||||
# 103| r0_23(void *) = Convert : r0_22
|
||||
# 103| v0_24(void) = Call : r0_21, r0_23
|
||||
# 103| mu0_25(unknown) = ^CallSideEffect : mu0_2
|
||||
# 104| v0_26(void) = NoOp :
|
||||
# 100| v0_27(void) = ReturnVoid :
|
||||
# 100| v0_28(void) = UnmodeledUse : mu*
|
||||
# 100| v0_29(void) = ExitFunction :
|
||||
|
||||
@@ -108,7 +108,7 @@ test.cpp:
|
||||
# 16| valnum = r0_17
|
||||
# 16| r0_18(glval<int>) = VariableAddress[global01] :
|
||||
# 16| valnum = r0_18
|
||||
# 16| r0_19(int) = Load : r0_18, mu0_2
|
||||
# 16| r0_19(int) = Load : r0_18, m0_1
|
||||
# 16| valnum = unique
|
||||
# 16| r0_20(int) = Add : r0_17, r0_19
|
||||
# 16| valnum = r0_20
|
||||
@@ -128,7 +128,7 @@ test.cpp:
|
||||
# 17| valnum = r0_17
|
||||
# 17| r0_28(glval<int>) = VariableAddress[global01] :
|
||||
# 17| valnum = r0_18
|
||||
# 17| r0_29(int) = Load : r0_28, mu0_2
|
||||
# 17| r0_29(int) = Load : r0_28, m0_1
|
||||
# 17| valnum = unique
|
||||
# 17| r0_30(int) = Add : r0_27, r0_29
|
||||
# 17| valnum = r0_30
|
||||
@@ -190,7 +190,7 @@ test.cpp:
|
||||
# 29| valnum = r0_17
|
||||
# 29| r0_18(glval<int>) = VariableAddress[global02] :
|
||||
# 29| valnum = r0_18
|
||||
# 29| r0_19(int) = Load : r0_18, mu0_2
|
||||
# 29| r0_19(int) = Load : r0_18, m0_1
|
||||
# 29| valnum = unique
|
||||
# 29| r0_20(int) = Add : r0_17, r0_19
|
||||
# 29| valnum = r0_20
|
||||
@@ -217,7 +217,7 @@ test.cpp:
|
||||
# 31| valnum = r0_17
|
||||
# 31| r0_32(glval<int>) = VariableAddress[global02] :
|
||||
# 31| valnum = r0_18
|
||||
# 31| r0_33(int) = Load : r0_32, mu0_2
|
||||
# 31| r0_33(int) = Load : r0_32, m0_26
|
||||
# 31| valnum = unique
|
||||
# 31| r0_34(int) = Add : r0_31, r0_33
|
||||
# 31| valnum = r0_34
|
||||
@@ -283,7 +283,7 @@ test.cpp:
|
||||
# 43| valnum = r0_19
|
||||
# 43| r0_20(glval<int>) = VariableAddress[global03] :
|
||||
# 43| valnum = r0_20
|
||||
# 43| r0_21(int) = Load : r0_20, mu0_2
|
||||
# 43| r0_21(int) = Load : r0_20, m0_1
|
||||
# 43| valnum = unique
|
||||
# 43| r0_22(int) = Add : r0_19, r0_21
|
||||
# 43| valnum = r0_22
|
||||
@@ -313,7 +313,7 @@ test.cpp:
|
||||
# 45| valnum = r0_19
|
||||
# 45| r0_35(glval<int>) = VariableAddress[global03] :
|
||||
# 45| valnum = r0_20
|
||||
# 45| r0_36(int) = Load : r0_35, mu0_2
|
||||
# 45| r0_36(int) = Load : r0_35, m0_29
|
||||
# 45| valnum = unique
|
||||
# 45| r0_37(int) = Add : r0_34, r0_36
|
||||
# 45| valnum = r0_37
|
||||
|
||||
Reference in New Issue
Block a user