diff --git a/mlsource/MLCompiler/BUILTINS.sml b/mlsource/MLCompiler/BUILTINS.sml index 13a92031..254d39f5 100644 --- a/mlsource/MLCompiler/BUILTINS.sml +++ b/mlsource/MLCompiler/BUILTINS.sml @@ -1,102 +1,106 @@ (* Signature for built-in functions - Copyright David C. J. Matthews 2016, 2018-9 + Copyright David C. J. Matthews 2016, 2018-20 This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) signature BUILTINS = sig datatype testConditions = TestEqual (* No TestNotEqual because that is always generated with "not" *) | TestLess | TestLessEqual | TestGreater | TestGreaterEqual | TestUnordered (* Reals only. *) datatype arithmeticOperations = ArithAdd | ArithSub | ArithMult | ArithQuot | ArithRem | ArithDiv | ArithMod datatype logicalOperations = LogicalAnd | LogicalOr | LogicalXor datatype shiftOperations = ShiftLeft | ShiftRightLogical (* Logical shift - zero added bits. *) | ShiftRightArithmetic (* Arithmetic shift - add the sign bit. *) datatype unaryOps = NotBoolean (* true => false; false => true - XOR *) | IsTaggedValue (* Test the tag bit. *) | MemoryCellLength (* Return the length of a memory cell (heap object) *) | MemoryCellFlags (* Return the flags byte of a memory cell (heap object) *) | ClearMutableFlag (* Remove the mutable flag from the flags byte *) | AtomicIncrement | AtomicDecrement | AtomicReset (* Set a value to (tagged) zero atomically. *) | LongWordToTagged (* Convert a LargeWord.word to a Word.word or FixedInt.int. *) | SignedToLongWord (* Convert a tagged value to a LargeWord with sign extension. *) | UnsignedToLongWord (* Convert a tagged value to a LargeWord without sign extension. *) | RealAbs of precision (* Set the sign bit of a real to positive. *) | RealNeg of precision (* Invert the sign bit of a real. *) | RealFixedInt of precision (* Convert an integer value into a real value. *) | FloatToDouble (* Convert a single precision floating point value to double precision. *) | DoubleToFloat of IEEEReal.rounding_mode option (* Convert a double precision floating point value to single precision. *) | RealToInt of precision * IEEEReal.rounding_mode (* Convert a double or float to a fixed precision int. *) | TouchAddress (* Ensures that the cell is reachable. *) and precision = PrecSingle | PrecDouble (* Single or double precision floating pt. *) and binaryOps = (* Compare two words and return the result. This is used for both word values (isSigned=false) and fixed precision integer (isSigned=true). - Tests for (in)equality can also be done on pointers in which case - this is pointer equality. *) + Values must be tagged and not pointers. *) WordComparison of { test: testConditions, isSigned: bool } (* Fixed precision int operations. These may raise Overflow. *) | FixedPrecisionArith of arithmeticOperations (* Arithmetic operations on word values. These do not raise Overflow. *) | WordArith of arithmeticOperations (* Load a word at a specific offset in a heap object. If this is immutable and the arguments are constants it can be folded at compile time since the result will never change. *) | WordLogical of logicalOperations (* Logical operations on words. *) | WordShift of shiftOperations (* Shift operations on words. *) (* Allocate a heap cell for byte data. The first argument is the number of words (not bytes) needed. The second argument is the "flags" byte which must include F_bytes and F_mutable. The new cell is not initialised. *) | AllocateByteMemory (* Operations on LargeWords. These are 32/64 bit values that are "boxed". *) | LargeWordComparison of testConditions | LargeWordArith of arithmeticOperations | LargeWordLogical of logicalOperations | LargeWordShift of shiftOperations | RealComparison of testConditions * precision | RealArith of arithmeticOperations * precision + (* Equality of values which could be pointers or tagged values. + At the lowest level this is the same as WordComparison but + if we try to use an indexed case there must be a check that the + values are tagged. *) + | PointerEq val unaryRepr: unaryOps -> string and binaryRepr: binaryOps -> string and testRepr: testConditions -> string and arithRepr: arithmeticOperations -> string end; diff --git a/mlsource/MLCompiler/CODETREESIG.ML b/mlsource/MLCompiler/CODETREESIG.ML index bb3dd98f..3da8ac22 100644 --- a/mlsource/MLCompiler/CODETREESIG.ML +++ b/mlsource/MLCompiler/CODETREESIG.ML @@ -1,159 +1,160 @@ (* - Copyright (c) 2012,13,15-18 David C.J. Matthews + Copyright (c) 2012,13,15-18, 20 David C.J. Matthews This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) signature CODETREESIG = sig type machineWord type codetree type pretty type codeBinding type level datatype argumentType = GeneralType | DoubleFloatType | SingleFloatType and loadStoreKind = LoadStoreMLWord of {isImmutable: bool} | LoadStoreMLByte of {isImmutable: bool} | LoadStoreC8 | LoadStoreC16 | LoadStoreC32 | LoadStoreC64 | LoadStoreCFloat | LoadStoreCDouble | LoadStoreUntaggedUnsigned and blockOpKind = BlockOpMove of {isByteMove: bool} | BlockOpEqualByte | BlockOpCompareByte structure BuiltIns: BUILTINS datatype arbPrecisionOps = ArbCompare of BuiltIns.testConditions | ArbArith of BuiltIns.arithmeticOperations val CodeTrue: codetree (* code for "true" *) val CodeFalse: codetree (* code for "false" *) val CodeZero: codetree (* code for 0, nil etc. *) val mkFunction: { body: codetree, argTypes:argumentType list, resultType: argumentType, name: string, closure: codetree list, numLocals: int } -> codetree val mkInlineFunction: { body: codetree, argTypes:argumentType list, resultType: argumentType, name: string, closure: codetree list, numLocals: int } -> codetree val mkCall: codetree * (codetree * argumentType) list * argumentType -> codetree val mkLoadLocal: int -> codetree and mkLoadArgument: int -> codetree and mkLoadClosure: int -> codetree val mkConst: machineWord -> codetree val mkInd: int * codetree -> codetree val mkVarField: int * codetree -> codetree val mkProc: codetree * int * string * codetree list * int -> codetree val mkInlproc: codetree * int * string * codetree list * int -> codetree val mkMacroProc: codetree * int * string * codetree list * int -> codetree val mkIf: codetree * codetree * codetree -> codetree val mkWhile: codetree * codetree -> codetree val mkEnv: codeBinding list * codetree -> codetree val mkStr: string -> codetree val mkTuple: codetree list -> codetree val mkDatatype: codetree list -> codetree val mkRaise: codetree -> codetree val mkCor: codetree * codetree -> codetree val mkCand: codetree * codetree -> codetree val mkHandle: codetree * codetree * int -> codetree val mkEval: codetree * codetree list -> codetree val identityFunction: string -> codetree val mkSetContainer: codetree * codetree * int -> codetree val mkTupleFromContainer: int * int -> codetree val mkTagTest: codetree * word * word -> codetree val mkBeginLoop: codetree * (int * codetree) list -> codetree val mkLoop: codetree list -> codetree val mkDec: int * codetree -> codeBinding val mkMutualDecs: (int * codetree) list -> codeBinding val mkNullDec: codetree -> codeBinding val mkContainer: int * int * codetree -> codeBinding - val mkNot: codetree -> codetree - val mkIsShort: codetree -> codetree - val mkEqualWord: codetree * codetree -> codetree - val mkEqualArbShort: codetree * codetree -> codetree - val equalWordFn: codetree + val mkNot: codetree -> codetree + val mkIsShort: codetree -> codetree + val mkEqualTaggedWord: codetree * codetree -> codetree + val mkEqualPointerOrWord: codetree * codetree -> codetree + val equalTaggedWordFn: codetree + val equalPointerOrWordFn: codetree val decSequenceWithFinalExp: codeBinding list -> codetree val pretty: codetree -> pretty val evalue: codetree -> machineWord option val genCode: codetree * Universal.universal list * int -> (unit -> codetree) (* Helper functions to build closure. *) val mkLoad: int * level * level -> codetree and mkLoadParam: int * level * level -> codetree val baseLevel: level val newLevel: level -> level val getClosure: level -> codetree list val multipleUses: codetree * (unit -> int) * level -> {load: level -> codetree, dec: codeBinding list} val mkUnary: BuiltIns.unaryOps * codetree -> codetree and mkBinary: BuiltIns.binaryOps * codetree * codetree -> codetree val mkUnaryFn: BuiltIns.unaryOps -> codetree and mkBinaryFn: BuiltIns.binaryOps -> codetree and mkArbitraryFn: arbPrecisionOps -> codetree val getCurrentThreadId: codetree and getCurrentThreadIdFn: codetree val mkAllocateWordMemory: codetree * codetree * codetree -> codetree and mkAllocateWordMemoryFn: codetree (* Load and store operations. At this level the first operand is the base address and the second is an index. *) val mkLoadOperation: loadStoreKind * codetree * codetree -> codetree val mkLoadOperationFn: loadStoreKind -> codetree val mkStoreOperation: loadStoreKind * codetree * codetree * codetree -> codetree val mkStoreOperationFn: loadStoreKind -> codetree val mkBlockOperation: {kind:blockOpKind, leftBase: codetree, rightBase: codetree, leftIndex: codetree, rightIndex: codetree, length: codetree} -> codetree val mkBlockOperationFn: blockOpKind -> codetree structure Foreign: FOREIGNCALLSIG structure Sharing: sig type machineWord = machineWord type codetree = codetree type pretty = pretty type argumentType=argumentType type codeBinding = codeBinding type level = level end end; diff --git a/mlsource/MLCompiler/CodeTree/BackendIntermediateCode.sml b/mlsource/MLCompiler/CodeTree/BackendIntermediateCode.sml index 8b2323fe..534b3364 100644 --- a/mlsource/MLCompiler/CodeTree/BackendIntermediateCode.sml +++ b/mlsource/MLCompiler/CodeTree/BackendIntermediateCode.sml @@ -1,726 +1,728 @@ (* - Copyright (c) 2012, 2016-19 David C.J. Matthews + Copyright (c) 2012, 2016-20 David C.J. Matthews This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) (* Intermediate code tree for the back end of the compiler. *) structure BackendIntermediateCode: BackendIntermediateCodeSig = struct open Address structure BuiltIns = struct datatype testConditions = TestEqual | TestLess | TestLessEqual | TestGreater | TestGreaterEqual | TestUnordered (* Reals only. *) datatype arithmeticOperations = ArithAdd | ArithSub | ArithMult | ArithQuot | ArithRem | ArithDiv | ArithMod datatype logicalOperations = LogicalAnd | LogicalOr | LogicalXor datatype shiftOperations = ShiftLeft | ShiftRightLogical | ShiftRightArithmetic datatype unaryOps = NotBoolean | IsTaggedValue | MemoryCellLength | MemoryCellFlags | ClearMutableFlag | AtomicIncrement | AtomicDecrement | AtomicReset | LongWordToTagged | SignedToLongWord | UnsignedToLongWord | RealAbs of precision | RealNeg of precision | RealFixedInt of precision | FloatToDouble | DoubleToFloat of IEEEReal.rounding_mode option | RealToInt of precision * IEEEReal.rounding_mode | TouchAddress and precision = PrecSingle | PrecDouble and binaryOps = WordComparison of { test: testConditions, isSigned: bool } | FixedPrecisionArith of arithmeticOperations | WordArith of arithmeticOperations | WordLogical of logicalOperations | WordShift of shiftOperations | AllocateByteMemory | LargeWordComparison of testConditions | LargeWordArith of arithmeticOperations | LargeWordLogical of logicalOperations | LargeWordShift of shiftOperations | RealComparison of testConditions * precision | RealArith of arithmeticOperations * precision + | PointerEq fun unaryRepr NotBoolean = "NotBoolean" | unaryRepr IsTaggedValue = "IsTaggedValue" | unaryRepr MemoryCellLength = "MemoryCellLength" | unaryRepr MemoryCellFlags = "MemoryCellFlags" | unaryRepr ClearMutableFlag = "ClearMutableFlag" | unaryRepr AtomicIncrement = "AtomicIncrement" | unaryRepr AtomicDecrement = "AtomicDecrement" | unaryRepr AtomicReset = "AtomicReset" | unaryRepr LongWordToTagged = "LongWordToTagged" | unaryRepr SignedToLongWord = "SignedToLongWord" | unaryRepr UnsignedToLongWord = "UnsignedToLongWord" | unaryRepr (RealAbs prec) = "RealAbs" ^ precRepr prec | unaryRepr (RealNeg prec) = "RealNeg" ^ precRepr prec | unaryRepr (RealFixedInt prec) = "RealFixedInt" ^ precRepr prec | unaryRepr FloatToDouble = "FloatToDouble" | unaryRepr (DoubleToFloat NONE) = "DoubleToFloat" | unaryRepr (DoubleToFloat (SOME mode)) = "DoubleToFloat" ^ rndModeRepr mode | unaryRepr (RealToInt (prec, mode)) = "RealToInt" ^ precRepr prec ^ rndModeRepr mode | unaryRepr TouchAddress = "TouchAddress" and binaryRepr (WordComparison{test, isSigned}) = "Test" ^ (testRepr test) ^ (if isSigned then "Signed" else "Unsigned") | binaryRepr (FixedPrecisionArith arithOp) = (arithRepr arithOp) ^ "Fixed" | binaryRepr (WordArith arithOp) = (arithRepr arithOp) ^ "Word" | binaryRepr (WordLogical logOp) = (logicRepr logOp) ^ "Word" | binaryRepr (WordShift shiftOp) = (shiftRepr shiftOp) ^ "Word" | binaryRepr AllocateByteMemory = "AllocateByteMemory" | binaryRepr (LargeWordComparison test) = "Test" ^ (testRepr test) ^ "LargeWord" | binaryRepr (LargeWordArith arithOp) = (arithRepr arithOp) ^ "LargeWord" | binaryRepr (LargeWordLogical logOp) = (logicRepr logOp) ^ "LargeWord" | binaryRepr (LargeWordShift shiftOp) = (shiftRepr shiftOp) ^ "LargeWord" | binaryRepr (RealComparison (test, prec)) = "Test" ^ testRepr test ^ precRepr prec | binaryRepr (RealArith (arithOp, prec)) = arithRepr arithOp ^ precRepr prec + | binaryRepr PointerEq = "PointerEq" and testRepr TestEqual = "Equal" | testRepr TestLess = "Less" | testRepr TestLessEqual = "LessEqual" | testRepr TestGreater = "Greater" | testRepr TestGreaterEqual = "GreaterEqual" | testRepr TestUnordered = "Unordered" and arithRepr ArithAdd = "Add" | arithRepr ArithSub = "Sub" | arithRepr ArithMult = "Mult" | arithRepr ArithQuot = "Quot" | arithRepr ArithRem = "Rem" | arithRepr ArithDiv = "Div" | arithRepr ArithMod = "Mod" and logicRepr LogicalAnd = "And" | logicRepr LogicalOr = "Or" | logicRepr LogicalXor = "Xor" and shiftRepr ShiftLeft = "Left" | shiftRepr ShiftRightLogical = "RightLogical" | shiftRepr ShiftRightArithmetic = "RightArithmetic" and precRepr PrecSingle = "Single" | precRepr PrecDouble = "Double" and rndModeRepr IEEEReal.TO_NEAREST = "Round" | rndModeRepr IEEEReal.TO_NEGINF = "Down" | rndModeRepr IEEEReal.TO_POSINF = "Up" | rndModeRepr IEEEReal.TO_ZERO = "Trunc" end datatype argumentType = GeneralType | DoubleFloatType | SingleFloatType datatype backendIC = BICNewenv of bicCodeBinding list * backendIC (* Set of bindings with an expression. *) | BICConstnt of machineWord * Universal.universal list (* Load a constant *) | BICExtract of bicLoadForm (* Get a local variable, an argument or a closure value *) | BICField of {base: backendIC, offset: int } (* Load a field from a tuple or record *) | BICEval of (* Evaluate a function with an argument list. *) { function: backendIC, argList: (backendIC * argumentType) list, resultType: argumentType } (* Built-in functions. *) | BICUnary of {oper: BuiltIns.unaryOps, arg1: backendIC} | BICBinary of {oper: BuiltIns.binaryOps, arg1: backendIC, arg2: backendIC} | BICArbitrary of {oper: BuiltIns.arithmeticOperations, shortCond: backendIC, arg1: backendIC, arg2: backendIC, longCall: backendIC} | BICLambda of bicLambdaForm (* Lambda expressions. *) | BICCond of backendIC * backendIC * backendIC (* If-then-else expression *) | BICCase of (* Case expressions *) { cases : backendIC option list, (* NONE means "jump to the default". *) test : backendIC, default : backendIC, isExhaustive: bool, firstIndex: word } | BICBeginLoop of (* Start of tail-recursive inline function. *) { loop: backendIC, arguments: (bicSimpleBinding * argumentType) list } | BICLoop of (backendIC * argumentType) list (* Jump back to start of tail-recursive function. *) | BICRaise of backendIC (* Raise an exception *) | BICHandle of (* Exception handler. *) { exp: backendIC, handler: backendIC, exPacketAddr: int } | BICTuple of backendIC list (* Tuple *) | BICSetContainer of (* Copy a tuple to a container. *) { container: backendIC, tuple: backendIC, filter: BoolVector.vector } | BICLoadContainer of {base: backendIC, offset: int } | BICTagTest of { test: backendIC, tag: word, maxTag: word } | BICLoadOperation of { kind: loadStoreKind, address: bicAddress } | BICStoreOperation of { kind: loadStoreKind, address: bicAddress, value: backendIC } | BICBlockOperation of { kind: blockOpKind, sourceLeft: bicAddress, destRight: bicAddress, length: backendIC } | BICGetThreadId | BICAllocateWordMemory of {numWords: backendIC, flags: backendIC, initial: backendIC} and bicCodeBinding = BICDeclar of bicSimpleBinding (* Make a local declaration or push an argument *) | BICRecDecs of { addr: int, lambda: bicLambdaForm } list (* Set of mutually recursive declarations. *) | BICNullBinding of backendIC (* Just evaluate the expression and discard the result. *) | BICDecContainer of { addr: int, size: int } (* Create a container for a tuple on the stack. *) and caseType = CaseWord (* Word or fixed-precision integer. *) | CaseTag of word and bicLoadForm = BICLoadLocal of int (* Local binding *) | BICLoadArgument of int (* Argument - 0 is first arg etc.*) | BICLoadClosure of int (* Closure - 0 is first closure item etc *) | BICLoadRecursive (* Recursive call *) and loadStoreKind = LoadStoreMLWord of {isImmutable: bool} (* Load/Store an ML word in an ML word cell. *) | LoadStoreMLByte of {isImmutable: bool} (* Load/Store a byte, tagging and untagging as appropriate, in an ML byte cell. *) | LoadStoreC8 (* Load/Store C values - The base address is a boxed SysWord.word value. *) | LoadStoreC16 | LoadStoreC32 | LoadStoreC64 | LoadStoreCFloat | LoadStoreCDouble | LoadStoreUntaggedUnsigned and blockOpKind = BlockOpMove of {isByteMove: bool} | BlockOpEqualByte | BlockOpCompareByte withtype bicSimpleBinding = { (* Declare a value or push an argument. *) value: backendIC, addr: int } and bicLambdaForm = { (* Lambda expressions. *) body : backendIC, name : string, closure : bicLoadForm list, argTypes : argumentType list, resultType : argumentType, localCount : int, heapClosure : bool } and bicAddress = (* Address form used in loads, store and block operations. The base is an ML address if this is to/from ML memory or a (boxed) SysWord.word if it is to/from C memory. The index is a value in units of the size of the item being loaded/stored and the offset is always in bytes. *) {base: backendIC, index: backendIC option, offset: word} structure CodeTags = struct open Universal val tupleTag: universal list list tag = tag() fun splitProps _ [] = (NONE, []) | splitProps tag (hd::tl) = if Universal.tagIs tag hd then (SOME hd, tl) else let val (p, l) = splitProps tag tl in (p, hd :: l) end fun mergeTupleProps(p, []) = p | mergeTupleProps([], p) = p | mergeTupleProps(m, n) = ( case (splitProps tupleTag m, splitProps tupleTag n) of ((SOME mp, ml), (SOME np, nl)) => let val mpl = Universal.tagProject tupleTag mp and npl = Universal.tagProject tupleTag np val merge = ListPair.mapEq mergeTupleProps (mpl, npl) in Universal.tagInject tupleTag merge :: (ml @ nl) end | _ => m @ n ) end fun loadStoreKindRepr(LoadStoreMLWord {isImmutable=true}) = "MLWordImmutable" | loadStoreKindRepr(LoadStoreMLWord {isImmutable=false}) = "MLWord" | loadStoreKindRepr(LoadStoreMLByte {isImmutable=true}) = "MLByteImmutable" | loadStoreKindRepr(LoadStoreMLByte {isImmutable=false}) = "MLByte" | loadStoreKindRepr LoadStoreC8 = "C8Bit" | loadStoreKindRepr LoadStoreC16 = "C16Bit" | loadStoreKindRepr LoadStoreC32 = "C32Bit" | loadStoreKindRepr LoadStoreC64 = "C64Bit" | loadStoreKindRepr LoadStoreCFloat = "CFloat" | loadStoreKindRepr LoadStoreCDouble = "CDouble" | loadStoreKindRepr LoadStoreUntaggedUnsigned = "MLWordUntagged" fun blockOpKindRepr (BlockOpMove{isByteMove=false}) = "MoveWord" | blockOpKindRepr (BlockOpMove{isByteMove=true}) = "MoveByte" | blockOpKindRepr BlockOpEqualByte = "EqualByte" | blockOpKindRepr BlockOpCompareByte = "CompareByte" open Pretty fun pList ([]: 'b list, _: string, _: 'b->pretty) = [] | pList ([h], _, disp) = [disp h] | pList (h::t, sep, disp) = PrettyBlock (0, false, [], [ disp h, PrettyBreak (0, 0), PrettyString sep ] ) :: PrettyBreak (1, 0) :: pList (t, sep, disp) fun pretty (pt : backendIC) : pretty = let fun printList(start, lst, sep) : pretty = PrettyBlock (1, true, [], PrettyString (start ^ "(") :: pList(lst, sep, pretty) @ [ PrettyBreak (0, 0), PrettyString (")") ] ) fun prettyArgType GeneralType = PrettyString "G" | prettyArgType DoubleFloatType = PrettyString "D" | prettyArgType SingleFloatType = PrettyString "F" fun prettyArg (c, t) = PrettyBlock(1, false, [], [pretty c, PrettyBreak (1, 0), prettyArgType t]) fun prettyArgs(start, lst, sep) : pretty = PrettyBlock (1, true, [], PrettyString (start ^ "(") :: pList(lst, sep, prettyArg) @ [ PrettyBreak (0, 0), PrettyString (")") ] ) fun prettyAddress({base, index, offset}: bicAddress): pretty = let in PrettyBlock (1, true, [], [ PrettyString "[", PrettyBreak (0, 3), pretty base, PrettyBreak (0, 0), PrettyString ",", PrettyBreak (1, 0), case index of NONE => PrettyString "-" | SOME i => pretty i, PrettyBreak (0, 0), PrettyString ",", PrettyBreak (1, 0), PrettyString(Word.toString offset), PrettyBreak (0, 0), PrettyString "]" ]) end in case pt of BICEval {function, argList, resultType} => let val prettyArgs = PrettyBlock (1, true, [], PrettyString ("$(") :: pList(argList, ",", prettyArg) @ [ PrettyBreak (0, 0), PrettyString (")") ] ) in PrettyBlock (3, false, [], [ pretty function, PrettyBreak(1, 0), prettyArgType resultType, PrettyBreak(1, 0), prettyArgs ] ) end | BICGetThreadId => PrettyString "GetThreadId" | BICUnary { oper, arg1 } => PrettyBlock (3, false, [], [ PrettyString(BuiltIns.unaryRepr oper), PrettyBreak(1, 0), printList("", [arg1], ",") ] ) | BICBinary { oper, arg1, arg2 } => PrettyBlock (3, false, [], [ PrettyString(BuiltIns.binaryRepr oper), PrettyBreak(1, 0), printList("", [arg1, arg2], ",") ] ) | BICArbitrary { oper, shortCond, arg1, arg2, longCall } => PrettyBlock (3, false, [], [ PrettyString(BuiltIns.arithRepr oper), PrettyBreak(1, 0), printList("", [shortCond, arg1, arg2, longCall], ",") ] ) | BICAllocateWordMemory { numWords, flags, initial } => PrettyBlock (3, false, [], [ PrettyString "AllocateWordMemory", PrettyBreak(1, 0), printList("", [numWords, flags, initial], ",") ] ) | BICExtract (BICLoadLocal addr) => let val str : string = concat ["LOCAL(", Int.toString addr, ")"] in PrettyString str end | BICExtract (BICLoadArgument addr) => let val str : string = concat ["PARAM(", Int.toString addr, ")"] in PrettyString str end | BICExtract (BICLoadClosure addr) => let val str : string = concat ["CLOS(", Int.toString addr, ")"] in PrettyString str end | BICExtract (BICLoadRecursive) => let val str : string = concat ["RECURSIVE(", ")"] in PrettyString str end | BICField {base, offset} => let val str = "INDIRECT(" ^ Int.toString offset ^ ", "; in PrettyBlock(0, false, [], [ PrettyString str, pretty base, PrettyString ")" ] ) end | BICLambda {body, name, closure, argTypes, heapClosure, resultType, localCount} => let fun prettyArgTypes [] = [] | prettyArgTypes [last] = [prettyArgType last] | prettyArgTypes (hd::tl) = prettyArgType hd :: PrettyBreak(1, 0) :: prettyArgTypes tl in PrettyBlock (1, true, [], [ PrettyString ("LAMBDA("), PrettyBreak (1, 0), PrettyString name, PrettyBreak (1, 0), PrettyString ( "CL=" ^ Bool.toString heapClosure), PrettyString (" LOCALS=" ^ Int.toString localCount), PrettyBreak(1, 0), PrettyBlock (1, false, [], PrettyString "ARGS=" :: prettyArgTypes argTypes), PrettyBreak(1, 0), PrettyBlock (1, false, [], [PrettyString "RES=", prettyArgType resultType]), printList (" CLOS=", map BICExtract closure, ","), PrettyBreak (1, 0), pretty body, PrettyString "){LAMBDA}" ] ) end | BICConstnt (w, _) => PrettyString (stringOfWord w) | BICCond (f, s, t) => PrettyBlock (1, true, [], [ PrettyString "IF(", pretty f, PrettyString ", ", PrettyBreak (0, 0), pretty s, PrettyString ", ", PrettyBreak (0, 0), pretty t, PrettyBreak (0, 0), PrettyString (")") ] ) | BICNewenv(decs, final) => PrettyBlock (1, true, [], PrettyString ("BLOCK" ^ "(") :: pList(decs, ";", prettyBinding) @ [ PrettyBreak (1, 0), pretty final, PrettyBreak (0, 0), PrettyString (")") ] ) | BICBeginLoop{loop=loopExp, arguments=args } => let fun prettyArg (c, t) = PrettyBlock(1, false, [], [prettySimpleBinding c, PrettyBreak (1, 0), prettyArgType t]) in PrettyBlock (3, false, [], [ PrettyBlock (1, true, [], PrettyString ("BEGINLOOP(") :: pList(args, ",", prettyArg) @ [ PrettyBreak (0, 0), PrettyString (")") ] ), PrettyBreak (0, 0), PrettyString "(", PrettyBreak (0, 0), pretty loopExp, PrettyBreak (0, 0), PrettyString ")" ] ) end | BICLoop ptl => prettyArgs("LOOP", ptl, ",") | BICRaise c => PrettyBlock (1, true, [], [ PrettyString "RAISE(", pretty c, PrettyBreak (0, 0), PrettyString (")") ] ) | BICHandle {exp, handler, exPacketAddr} => PrettyBlock (3, false, [], [ PrettyString "HANDLE(", pretty exp, PrettyString ("WITH exid=" ^ Int.toString exPacketAddr), PrettyBreak (1, 0), pretty handler, PrettyString ")" ] ) | BICCase {cases, test, default, isExhaustive, firstIndex, ...} => PrettyBlock (1, true, [], PrettyString "CASE (" :: pretty test :: PrettyBreak (1, 0) :: PrettyString ("( from " ^ Word.toString firstIndex ^ (if isExhaustive then " exhaustive" else "")) :: PrettyBreak (1, 0) :: pList(cases, ",", fn (SOME exp) => PrettyBlock (1, true, [], [ PrettyString "=>", PrettyBreak (1, 0), pretty exp ]) | NONE => PrettyString "=> default" ) @ [ PrettyBreak (1, 0), PrettyBlock (1, false, [], [ PrettyString "ELSE:", PrettyBreak (1, 0), pretty default ] ), PrettyBreak (1, 0), PrettyString (") {"^"CASE"^"}") ] ) | BICTuple ptl => printList("RECCONSTR", ptl, ",") | BICSetContainer{container, tuple, filter} => let val source = BoolVector.length filter val dest = BoolVector.foldl(fn (true, n) => n+1 | (false, n) => n) 0 filter in PrettyBlock (3, false, [], [ PrettyString (concat["SETCONTAINER(", Int.toString dest, "/", Int.toString source, ", "]), pretty container, PrettyBreak (0, 0), PrettyString ",", PrettyBreak (1, 0), pretty tuple, PrettyBreak (0, 0), PrettyString ")" ] ) end | BICLoadContainer {base, offset} => let val str = "INDIRECTCONTAINER(" ^ Int.toString offset ^ ", "; in PrettyBlock(0, false, [], [ PrettyString str, pretty base, PrettyString ")" ] ) end | BICTagTest { test, tag, maxTag } => PrettyBlock (3, false, [], [ PrettyString (concat["TAGTEST(", Word.toString tag, ", ", Word.toString maxTag, ","]), PrettyBreak (1, 0), pretty test, PrettyBreak (0, 0), PrettyString ")" ] ) | BICLoadOperation{ kind, address } => PrettyBlock (3, false, [], [ PrettyString("Load" ^ loadStoreKindRepr kind), PrettyBreak (1, 0), prettyAddress address ] ) | BICStoreOperation{ kind, address, value } => PrettyBlock (3, false, [], [ PrettyString("Store" ^ loadStoreKindRepr kind), PrettyBreak (1, 0), prettyAddress address, PrettyBreak (1, 0), PrettyString "<=", PrettyBreak (1, 0), pretty value ] ) | BICBlockOperation{ kind, sourceLeft, destRight, length } => PrettyBlock (3, false, [], [ PrettyString(blockOpKindRepr kind ^ "("), PrettyBreak (1, 0), prettyAddress sourceLeft, PrettyBreak (1, 0), PrettyString ",", prettyAddress destRight, PrettyBreak (1, 0), PrettyString ",", pretty length, PrettyBreak (1, 0), PrettyString ")" ] ) (* That list should be exhaustive! *) end (* pretty *) and prettyBinding(BICDeclar dec) = prettySimpleBinding dec | prettyBinding(BICRecDecs ptl) = let fun prettyRDec {lambda, addr} = PrettyBlock (1, false, [], [ PrettyString (concat ["DECL #", Int.toString addr, "="]), PrettyBreak (1, 0), pretty(BICLambda lambda) ] ) in PrettyBlock (1, true, [], PrettyString ("MUTUAL" ^ "(") :: pList(ptl, " AND ", prettyRDec) @ [ PrettyBreak (0, 0), PrettyString (")") ] ) end | prettyBinding(BICNullBinding c) = pretty c | prettyBinding(BICDecContainer{addr, size}) = PrettyString (concat ["CONTAINER #", Int.toString addr, "=", Int.toString size]) and prettySimpleBinding{value, addr} = PrettyBlock (1, false, [], [ PrettyString (concat ["DECL #", Int.toString addr, "="]), PrettyBreak (1, 0), pretty value ] ) structure Sharing = struct type backendIC = backendIC and bicLoadForm = bicLoadForm and caseType = caseType and pretty = pretty and argumentType = argumentType and bicCodeBinding = bicCodeBinding and bicSimpleBinding = bicSimpleBinding and loadStoreKind = loadStoreKind and blockOpKind = blockOpKind and unaryOps = BuiltIns.unaryOps and binaryOps = BuiltIns.binaryOps and testConditions = BuiltIns.testConditions and arithmeticOperations = BuiltIns.arithmeticOperations end end; diff --git a/mlsource/MLCompiler/CodeTree/ByteCode/INTGCODE.ML b/mlsource/MLCompiler/CodeTree/ByteCode/INTGCODE.ML index 15646f36..fe5da4e9 100644 --- a/mlsource/MLCompiler/CodeTree/ByteCode/INTGCODE.ML +++ b/mlsource/MLCompiler/CodeTree/ByteCode/INTGCODE.ML @@ -1,1157 +1,1159 @@ (* Copyright (c) 2000 Cambridge University Technical Services Limited Further development copyright David C.J. Matthews 2016-18,2020 This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) (* Title: Generate interpretable code for Poly system from the code tree. Author: Dave Matthews, Cambridge University Computer Laboratory Copyright Cambridge University 1985 *) (* This generates byte-code that is interpreted by the run-time system. It is now used as a fall-back to allow Poly/ML to run on non-X86 architectures. Early versions were used as a porting aid while a native code-generator was being developed and the "enter-int" instructions that were needed for that have been retained although they no longer actually generate code. *) functor INTGCODE ( structure CODECONS : INTCODECONSSIG structure BACKENDTREE: BackendIntermediateCodeSig structure CODE_ARRAY: CODEARRAYSIG sharing CODECONS.Sharing = BACKENDTREE.Sharing = CODE_ARRAY.Sharing ) : GENCODESIG = struct open CODECONS open Address open BACKENDTREE open Misc open CODE_ARRAY val word0 = toMachineWord 0; val DummyValue : machineWord = word0; (* used as result of "raise e" etc. *) type caseForm = { cases : (backendIC * word) list, test : backendIC, caseType: caseType, default : backendIC } (* Where the result, if any, should go *) datatype whereto = NoResult (* discard result *) | ToStack (* Need a result but it can stay on the pseudo-stack *); (* Are we at the end of the function. *) datatype tail = EndOfProc | NotEnd (* Code generate a function or global declaration *) fun codegen (pt, cvec, resultClosure, numOfArgs, localCount, parameters) = let datatype decEntry = StackAddr of int | Empty val decVec = Array.array (localCount, Empty) (* Count of number of items on the stack. *) val realstackptr = ref 1 (* The closure ptr is already there *) (* Maximum size of the stack. *) val maxStack = ref 1 (* Exited - set to true if we have jumped out. *) val exited = ref false; (* Push a value onto the stack. *) fun incsp () = ( realstackptr := !realstackptr + 1; if !realstackptr > !maxStack then maxStack := !realstackptr else () ) (* An entry has been removed from the stack. *) fun decsp () = realstackptr := !realstackptr - 1; fun pushLocalStackValue addr = ( genLocal(!realstackptr + addr, cvec); incsp() ) (* Loads a local, argument or closure value; translating local stack addresses to real stack offsets. *) fun locaddr(BICLoadArgument locn) = pushLocalStackValue (numOfArgs-locn) | locaddr(BICLoadLocal locn) = ( (* positive address - on the stack. *) case Array.sub (decVec, locn) of StackAddr n => pushLocalStackValue (~ n) | _ => (* Should be on the stack, not a function. *) raise InternalError "locaddr: bad stack address" ) | locaddr(BICLoadClosure locn) = (* closure-pointer relative *) ( pushLocalStackValue ~1; (* The closure itself. *) genIndirect (locn+1, cvec) (* The value in the closure. +1 because first item is code addr. *) ) | locaddr BICLoadRecursive = pushLocalStackValue ~1 (* The closure itself - first value on the stack. *) (* generates code from the tree *) fun gencde (pt : backendIC, whereto : whereto, tailKind : tail, loopAddr) : unit = let (* Save the stack pointer value here. We may want to reset the stack. *) val oldsp = !realstackptr; (* Load the address and index value for byte operations. For ML memory operations the base is the address of an ML heap cell whereas for C operations it is a large-word box containing an address in C memory. That doesn't affect this code but the interpreter has to deal with these differently. *) fun genByteAddress{base, index, offset} = ( gencde (base, ToStack, NotEnd, loopAddr); (* Because the index and offset are both byte counts we can just add them if we need both. *) case (index, offset) of (NONE, offset) => (pushConst (toMachineWord offset, cvec); incsp()) | (SOME indexVal, 0w0) => gencde (indexVal, ToStack, NotEnd, loopAddr) | (SOME indexVal, offset) => ( gencde (indexVal, ToStack, NotEnd, loopAddr); pushConst (toMachineWord offset, cvec); genOpcode(opcode_wordAdd, cvec) ) ) (* Load the address, index value and offset for non-byte operations. Because the offset has already been scaled by the size of the operand we have to load the index and offset separately. *) fun genNonByteAddress{base, index, offset} = ( gencde (base, ToStack, NotEnd, loopAddr); case index of NONE => (pushConst (toMachineWord 0, cvec); incsp()) | SOME indexVal => gencde (indexVal, ToStack, NotEnd, loopAddr); pushConst (toMachineWord offset, cvec); incsp() ) val () = case pt of BICEval evl => genEval (evl, tailKind) | BICExtract ext => (* This may just be being used to discard a value which isn't used on this branch. *) if whereto = NoResult then () else locaddr ext | BICField {base, offset} => (gencde (base, ToStack, NotEnd, loopAddr); genIndirect (offset, cvec)) | BICLoadContainer {base, offset} => (gencde (base, ToStack, NotEnd, loopAddr); genIndirect (offset, cvec)) | BICLambda lam => genProc (lam, false, fn () => ()) | BICConstnt(w, _) => let val () = pushConst (w, cvec); in incsp () end | BICCond (testPart, thenPart, elsePart) => genCond (testPart, thenPart, elsePart, whereto, tailKind, loopAddr) | BICNewenv(decls, exp) => let (* Processes a list of entries. *) (* Mutually recursive declarations. May be either lambdas or constants. Recurse down the list pushing the addresses of the closure vectors, then unwind the recursion and fill them in. *) fun genMutualDecs [] = () | genMutualDecs ({lambda, addr, ...} :: otherDecs) = genProc (lambda, true, fn() => ( Array.update (decVec, addr, StackAddr (! realstackptr)); genMutualDecs (otherDecs) )) fun codeDecls(BICRecDecs dl) = genMutualDecs dl | codeDecls(BICDecContainer{size, addr}) = ( (* If this is a container we have to process it here otherwise it will be removed in the stack adjustment code. *) genContainer(size, cvec); (* Push the address of this container. *) realstackptr := !realstackptr + size + 1; (* Pushes N words plus the address. *) Array.update (decVec, addr, StackAddr(!realstackptr)) ) | codeDecls(BICDeclar{value, addr, ...}) = ( gencde (value, ToStack, NotEnd, loopAddr); Array.update (decVec, addr, StackAddr(!realstackptr)) ) | codeDecls(BICNullBinding exp) = gencde (exp, NoResult, NotEnd, loopAddr) in List.app codeDecls decls; gencde (exp, whereto, tailKind, loopAddr) end | BICBeginLoop {loop=body, arguments} => (* Execute the body which will contain at least one Loop instruction. There will also be path(s) which don't contain Loops and these will drop through. *) let val args = List.map #1 arguments (* Evaluate each of the arguments, pushing the result onto the stack. *) fun genLoopArg ({addr, value, ...}) = ( gencde (value, ToStack, NotEnd, loopAddr); Array.update (decVec, addr, StackAddr (!realstackptr)); !realstackptr (* Return the posn on the stack. *) ) val argIndexList = map genLoopArg args; val startSp = ! realstackptr; (* Remember the current top of stack. *) val startLoop = createLabel () val () = setLabel(startLoop, cvec) (* Start of loop *) in (* Process the body, passing the jump-back address down for the Loop instruction(s). *) gencde (body, whereto, tailKind, SOME(startLoop, startSp, argIndexList)) (* Leave the arguments on the stack. They can be cleared later if needed. *) end | BICLoop argList => (* Jump back to the enclosing BeginLoop. *) let val (startLoop, startSp, argIndexList) = case loopAddr of SOME l => l | NONE => raise InternalError "No BeginLoop for Loop instr" (* Evaluate the arguments. First push them to the stack because evaluating an argument may depend on the current value of others. Only when we've evaluated all of them can we overwrite the original argument positions. *) fun loadArgs ([], []) = !realstackptr - startSp (* The offset of all the args. *) | loadArgs (arg:: argList, _ :: argIndexList) = let (* Evaluate all the arguments. *) val () = gencde (arg, ToStack, NotEnd, NONE); val argOffset = loadArgs(argList, argIndexList); in genSetStackVal(argOffset, cvec); (* Copy the arg over. *) decsp(); (* The argument has now been popped. *) argOffset end | loadArgs _ = raise InternalError "loadArgs: Mismatched arguments"; val _: int = loadArgs(List.map #1 argList, argIndexList) in if !realstackptr <> startSp then resetStack (!realstackptr - startSp, false, cvec) (* Remove any local variables. *) else (); (* Jump back to the start of the loop. *) putBranchInstruction(Jump, startLoop, cvec) end | BICRaise exp => let val () = gencde (exp, ToStack, NotEnd, loopAddr) val () = genRaiseEx cvec; in exited := true end | BICHandle {exp, handler, exPacketAddr} => let (* Save old handler *) val () = genPushHandler cvec val () = incsp () val handlerLabel = createLabel() val () = putBranchInstruction (SetHandler, handlerLabel, cvec) val () = incsp() (* Code generate the body; "NotEnd" because we have to come back to remove the handler; "ToStack" because delHandler needs a result to carry down. *) val () = gencde (exp, ToStack, NotEnd, loopAddr) (* Now get out of the handler and restore the old one. *) val () = genOpcode(opcode_deleteHandler, cvec) val skipHandler = createLabel() val () = putBranchInstruction (Jump, skipHandler, cvec) (* Now process the handler itself. First we have to reset the stack. Note that we have to use "ToStack" again to be consistent with the stack-handling in the body-part. If we actually wanted "NoResult", the stack adjustment code at the end of gencde will take care of this. This means that I don't want to do any clever "end-of-function" optimisation either. SPF 6/1/97 *) val () = realstackptr := oldsp val () = exited := false val () = setLabel (handlerLabel, cvec) (* If we were executing machine code we must re-enter the interpreter. *) val () = genEnterIntCatch cvec (* Push the exception packet and set the address. *) val () = genLdexc cvec val () = incsp () val () = Array.update (decVec, exPacketAddr, StackAddr(!realstackptr)) val () = gencde (handler, ToStack, NotEnd, loopAddr) (* Have to remove the exception packet. *) val () = resetStack(1, true, cvec) val () = decsp() (* Finally fix-up the jump around the handler *) val () = setLabel (skipHandler, cvec) in exited := false end | BICCase ({cases, test, default, firstIndex, ...}) => let val () = gencde (test, ToStack, NotEnd, loopAddr) (* Label to jump to at the end of each case. *) val exitJump = createLabel() val () = if firstIndex = 0w0 then () else ( (* Subtract lower limit. Don't check for overflow. Instead allow large value to wrap around and check in "case" instruction. *) pushConst(toMachineWord firstIndex, cvec); genOpcode(opcode_wordSub, cvec) ) (* Generate the case instruction followed by the table of jumps. *) val nCases = List.length cases val caseLabels = genCase (nCases, cvec) val () = decsp () (* The default case, if any, follows the case statement. *) (* If we have a jump to the default set it to jump here. *) local fun fixDefault(NONE, defCase) = setLabel(defCase, cvec) | fixDefault(SOME _, _) = () in val () = ListPair.appEq fixDefault (cases, caseLabels) end val () = gencde (default, whereto, tailKind, loopAddr); val () = exited := false; fun genCases(SOME body, label) = ( (* First exit from the previous case or the default if this is the first. *) if !exited then () else putBranchInstruction(Jump, exitJump, cvec); (* Remove the result - the last case will leave it. *) case whereto of ToStack => decsp () | NoResult => (); (* Fix up the jump to come here. *) setLabel(label, cvec); exited := false; gencde (body, whereto, tailKind, loopAddr) ) | genCases(NONE, _) = () val () = ListPair.appEq genCases (cases, caseLabels) (* Finally set the exit jump to come here. *) val () = setLabel (exitJump, cvec) in exited := false end | BICTuple recList => let val size = List.length recList in (* Move the fields into the vector. *) List.app(fn v => gencde (v, ToStack, NotEnd, loopAddr)) recList; genTuple (size, cvec); realstackptr := !realstackptr - (size - 1) end | BICSetContainer{container, tuple, filter} => (* Copy the contents of a tuple into a container. If the tuple is a Tuple instruction we can avoid generating the tuple and then unpacking it and simply copy the fields that make up the tuple directly into the container. *) ( case tuple of BICTuple cl => (* Simply set the container from the values. *) let (* Load the address of the container. *) val _ = gencde (container, ToStack, NotEnd, loopAddr); fun setValues([], _, _) = () | setValues(v::tl, sourceOffset, destOffset) = if sourceOffset < BoolVector.length filter andalso BoolVector.sub(filter, sourceOffset) then ( gencde (v, ToStack, NotEnd, loopAddr); (* Move the entry into the container. This instruction pops the value to be moved but not the destination. *) genMoveToVec(destOffset, cvec); decsp(); setValues(tl, sourceOffset+1, destOffset+1) ) else setValues(tl, sourceOffset+1, destOffset) in setValues(cl, 0, 0) (* The container address is still on the stack. *) end | _ => let (* General case. *) (* First the target tuple, then the container. *) val () = gencde (tuple, ToStack, NotEnd, loopAddr) val () = gencde (container, ToStack, NotEnd, loopAddr) val last = BoolVector.foldli(fn (i, true, _) => i | (_, false, n) => n) ~1 filter fun copy (sourceOffset, destOffset) = if BoolVector.sub(filter, sourceOffset) then ( (* Duplicate the tuple address . *) genLocal(1, cvec); genIndirect(sourceOffset, cvec); genMoveToVec(destOffset, cvec); if sourceOffset = last then () else copy (sourceOffset+1, destOffset+1) ) else copy(sourceOffset+1, destOffset) in copy (0, 0) (* The container and tuple addresses are still on the stack. *) end ) | BICTagTest { test, tag, ... } => ( (* Convert this into a simple equality function. *) gencde (test, ToStack, NotEnd, loopAddr); pushConst (toMachineWord tag, cvec); genOpcode(opcode_equalWord, cvec) ) | BICGetThreadId => ( genOpcode(opcode_getThreadId, cvec); incsp() ) | BICUnary { oper, arg1 } => let open BuiltIns val () = gencde (arg1, ToStack, NotEnd, loopAddr) in case oper of NotBoolean => genOpcode(opcode_notBoolean, cvec) | IsTaggedValue => genOpcode(opcode_isTagged, cvec) | MemoryCellLength => genOpcode(opcode_cellLength, cvec) | MemoryCellFlags => genOpcode(opcode_cellFlags, cvec) | ClearMutableFlag => genOpcode(opcode_clearMutable, cvec) | AtomicIncrement => genOpcode(opcode_atomicIncr, cvec) | AtomicDecrement => genOpcode(opcode_atomicDecr, cvec) | AtomicReset => genOpcode(opcode_atomicReset, cvec) | LongWordToTagged => genOpcode(opcode_longWToTagged, cvec) | SignedToLongWord => genOpcode(opcode_signedToLongW, cvec) | UnsignedToLongWord => genOpcode(opcode_unsignedToLongW, cvec) | RealAbs PrecDouble => genOpcode(opcode_realAbs, cvec) | RealNeg PrecDouble => genOpcode(opcode_realNeg, cvec) | RealFixedInt PrecDouble => genOpcode(opcode_fixedIntToReal, cvec) | RealAbs PrecSingle => genOpcode(opcode_floatAbs, cvec) | RealNeg PrecSingle => genOpcode(opcode_floatNeg, cvec) | RealFixedInt PrecSingle => genOpcode(opcode_fixedIntToFloat, cvec) | FloatToDouble => genOpcode(opcode_floatToReal, cvec) | DoubleToFloat rnding => genDoubleToFloat(rnding, cvec) | RealToInt (PrecDouble, rnding) => genRealToInt(rnding, cvec) | RealToInt (PrecSingle, rnding) => genFloatToInt(rnding, cvec) | TouchAddress => resetStack(1, false, cvec) (* Discard this *) end | BICBinary { oper, arg1, arg2 } => let open BuiltIns val () = gencde (arg1, ToStack, NotEnd, loopAddr) val () = gencde (arg2, ToStack, NotEnd, loopAddr) in case oper of WordComparison{test=TestEqual, ...} => genOpcode(opcode_equalWord, cvec) | WordComparison{test=TestLess, isSigned=true} => genOpcode(opcode_lessSigned, cvec) | WordComparison{test=TestLessEqual, isSigned=true} => genOpcode(opcode_lessEqSigned, cvec) | WordComparison{test=TestGreater, isSigned=true} => genOpcode(opcode_greaterSigned, cvec) | WordComparison{test=TestGreaterEqual, isSigned=true} => genOpcode(opcode_greaterEqSigned, cvec) | WordComparison{test=TestLess, isSigned=false} => genOpcode(opcode_lessUnsigned, cvec) | WordComparison{test=TestLessEqual, isSigned=false} => genOpcode(opcode_lessEqUnsigned, cvec) | WordComparison{test=TestGreater, isSigned=false} => genOpcode(opcode_greaterUnsigned, cvec) | WordComparison{test=TestGreaterEqual, isSigned=false} => genOpcode(opcode_greaterEqUnsigned, cvec) | WordComparison{test=TestUnordered, ...} => raise InternalError "WordComparison: TestUnordered" + | PointerEq => genOpcode(opcode_equalWord, cvec) + | FixedPrecisionArith ArithAdd => genOpcode(opcode_fixedAdd, cvec) | FixedPrecisionArith ArithSub => genOpcode(opcode_fixedSub, cvec) | FixedPrecisionArith ArithMult => genOpcode(opcode_fixedMult, cvec) | FixedPrecisionArith ArithQuot => genOpcode(opcode_fixedQuot, cvec) | FixedPrecisionArith ArithRem => genOpcode(opcode_fixedRem, cvec) | FixedPrecisionArith ArithDiv => raise InternalError "TODO: FixedPrecisionArith ArithDiv" | FixedPrecisionArith ArithMod => raise InternalError "TODO: FixedPrecisionArith ArithMod" | WordArith ArithAdd => genOpcode(opcode_wordAdd, cvec) | WordArith ArithSub => genOpcode(opcode_wordSub, cvec) | WordArith ArithMult => genOpcode(opcode_wordMult, cvec) | WordArith ArithDiv => genOpcode(opcode_wordDiv, cvec) | WordArith ArithMod => genOpcode(opcode_wordMod, cvec) | WordArith _ => raise InternalError "WordArith - unimplemented instruction" | WordLogical LogicalAnd => genOpcode(opcode_wordAnd, cvec) | WordLogical LogicalOr => genOpcode(opcode_wordOr, cvec) | WordLogical LogicalXor => genOpcode(opcode_wordXor, cvec) | WordShift ShiftLeft => genOpcode(opcode_wordShiftLeft, cvec) | WordShift ShiftRightLogical => genOpcode(opcode_wordShiftRLog, cvec) | WordShift ShiftRightArithmetic => genOpcode(opcode_wordShiftRArith, cvec) | AllocateByteMemory => genOpcode(opcode_allocByteMem, cvec) | LargeWordComparison TestEqual => genOpcode(opcode_lgWordEqual, cvec) | LargeWordComparison TestLess => genOpcode(opcode_lgWordLess, cvec) | LargeWordComparison TestLessEqual => genOpcode(opcode_lgWordLessEq, cvec) | LargeWordComparison TestGreater => genOpcode(opcode_lgWordGreater, cvec) | LargeWordComparison TestGreaterEqual => genOpcode(opcode_lgWordGreaterEq, cvec) | LargeWordComparison TestUnordered => raise InternalError "LargeWordComparison: TestUnordered" | LargeWordArith ArithAdd => genOpcode(opcode_lgWordAdd, cvec) | LargeWordArith ArithSub => genOpcode(opcode_lgWordSub, cvec) | LargeWordArith ArithMult => genOpcode(opcode_lgWordMult, cvec) | LargeWordArith ArithDiv => genOpcode(opcode_lgWordDiv, cvec) | LargeWordArith ArithMod => genOpcode(opcode_lgWordMod, cvec) | LargeWordArith _ => raise InternalError "LargeWordArith - unimplemented instruction" | LargeWordLogical LogicalAnd => genOpcode(opcode_lgWordAnd, cvec) | LargeWordLogical LogicalOr => genOpcode(opcode_lgWordOr, cvec) | LargeWordLogical LogicalXor => genOpcode(opcode_lgWordXor, cvec) | LargeWordShift ShiftLeft => genOpcode(opcode_lgWordShiftLeft, cvec) | LargeWordShift ShiftRightLogical => genOpcode(opcode_lgWordShiftRLog, cvec) | LargeWordShift ShiftRightArithmetic => genOpcode(opcode_lgWordShiftRArith, cvec) | RealComparison (TestEqual, PrecDouble) => genOpcode(opcode_realEqual, cvec) | RealComparison (TestLess, PrecDouble) => genOpcode(opcode_realLess, cvec) | RealComparison (TestLessEqual, PrecDouble) => genOpcode(opcode_realLessEq, cvec) | RealComparison (TestGreater, PrecDouble) => genOpcode(opcode_realGreater, cvec) | RealComparison (TestGreaterEqual, PrecDouble) => genOpcode(opcode_realGreaterEq, cvec) | RealComparison (TestUnordered, PrecDouble) => genOpcode(opcode_realUnordered, cvec) | RealComparison (TestEqual, PrecSingle) => genOpcode(opcode_floatEqual, cvec) | RealComparison (TestLess, PrecSingle) => genOpcode(opcode_floatLess, cvec) | RealComparison (TestLessEqual, PrecSingle) => genOpcode(opcode_floatLessEq, cvec) | RealComparison (TestGreater, PrecSingle) => genOpcode(opcode_floatGreater, cvec) | RealComparison (TestGreaterEqual, PrecSingle) => genOpcode(opcode_floatGreaterEq, cvec) | RealComparison (TestUnordered, PrecSingle) => genOpcode(opcode_floatUnordered, cvec) | RealArith (ArithAdd, PrecDouble) => genOpcode(opcode_realAdd, cvec) | RealArith (ArithSub, PrecDouble) => genOpcode(opcode_realSub, cvec) | RealArith (ArithMult, PrecDouble) => genOpcode(opcode_realMult, cvec) | RealArith (ArithDiv, PrecDouble) => genOpcode(opcode_realDiv, cvec) | RealArith (ArithAdd, PrecSingle) => genOpcode(opcode_floatAdd, cvec) | RealArith (ArithSub, PrecSingle) => genOpcode(opcode_floatSub, cvec) | RealArith (ArithMult, PrecSingle) => genOpcode(opcode_floatMult, cvec) | RealArith (ArithDiv, PrecSingle) => genOpcode(opcode_floatDiv, cvec) | RealArith _ => raise InternalError "RealArith - unimplemented instruction" ; decsp() (* Removes one item from the stack. *) end | BICAllocateWordMemory {numWords as BICConstnt(length, _), flags as BICConstnt(flagByte, _), initial } => if isShort length andalso toShort length = 0w1 andalso isShort flagByte andalso toShort flagByte = 0wx40 then (* This is a very common case. *) ( gencde (initial, ToStack, NotEnd, loopAddr); genOpcode(opcode_alloc_ref, cvec) ) else let val () = gencde (numWords, ToStack, NotEnd, loopAddr) val () = gencde (flags, ToStack, NotEnd, loopAddr) val () = gencde (initial, ToStack, NotEnd, loopAddr) in genOpcode(opcode_allocWordMemory, cvec); decsp(); decsp() end | BICAllocateWordMemory { numWords, flags, initial } => let val () = gencde (numWords, ToStack, NotEnd, loopAddr) val () = gencde (flags, ToStack, NotEnd, loopAddr) val () = gencde (initial, ToStack, NotEnd, loopAddr) in genOpcode(opcode_allocWordMemory, cvec); decsp(); decsp() end | BICLoadOperation { kind=LoadStoreMLWord _, address={base, index=NONE, offset}} => ( (* If the index is a constant, frequently zero, we can use indirection. The offset is a byte count so has to be divided by the word size but it should always be an exact multiple. *) gencde (base, ToStack, NotEnd, loopAddr); offset mod wordSize = 0w0 orelse raise InternalError "gencde: BICLoadOperation - not word multiple"; genIndirect (Word.toInt(offset div wordSize), cvec) ) | BICLoadOperation { kind=LoadStoreMLWord _, address={base, index=SOME indexVal, offset}} => let (* Variable index. *) val () = gencde (base, ToStack, NotEnd, loopAddr) val () = gencde (indexVal, ToStack, NotEnd, loopAddr) val () = (pushConst (toMachineWord offset, cvec); incsp()) in genOpcode(opcode_loadMLWord, cvec); decsp(); decsp() end | BICLoadOperation { kind=LoadStoreMLByte _, address} => ( genByteAddress address; genOpcode(opcode_loadMLByte, cvec); decsp() ) | BICLoadOperation { kind=LoadStoreC8, address} => ( genByteAddress address; genOpcode(opcode_loadC8, cvec); decsp() ) | BICLoadOperation { kind=LoadStoreC16, address} => ( genNonByteAddress address; genOpcode(opcode_loadC16, cvec); decsp(); decsp() ) | BICLoadOperation { kind=LoadStoreC32, address} => ( genNonByteAddress address; genOpcode(opcode_loadC32, cvec); decsp(); decsp() ) | BICLoadOperation { kind=LoadStoreC64, address} => ( wordSize = 0w8 orelse raise InternalError "LoadStoreC64 but not 64-bit mode"; genNonByteAddress address; genOpcode(opcode_loadC64, cvec); decsp(); decsp() ) | BICLoadOperation { kind=LoadStoreCFloat, address} => ( genNonByteAddress address; genOpcode(opcode_loadCFloat, cvec); decsp(); decsp() ) | BICLoadOperation { kind=LoadStoreCDouble, address} => ( genNonByteAddress address; genOpcode(opcode_loadCDouble, cvec); decsp(); decsp() ) | BICLoadOperation { kind=LoadStoreUntaggedUnsigned, address} => ( genNonByteAddress address; genOpcode(opcode_loadUntagged, cvec); decsp(); decsp() ) | BICStoreOperation { kind=LoadStoreMLWord _, address={base, index=NONE, offset}, value } => let (* No index. We could almost use move_to_vec here except that it leaves the destination address on the stack instead of replacing it with "unit". *) val () = gencde (base, ToStack, NotEnd, loopAddr) val () = pushConst (toMachineWord 0, cvec) val () = incsp() val () = pushConst (toMachineWord offset, cvec) val () = incsp() val () = gencde (value, ToStack, NotEnd, loopAddr) in genOpcode(opcode_storeMLWord, cvec); decsp(); decsp(); decsp() end | BICStoreOperation { kind=LoadStoreMLWord _, address={base, index=SOME indexVal, offset}, value } => let (* Variable index *) val () = gencde (base, ToStack, NotEnd, loopAddr) val () = gencde (indexVal, ToStack, NotEnd, loopAddr) val () = pushConst (toMachineWord offset, cvec) val () = incsp() val () = gencde (value, ToStack, NotEnd, loopAddr) in genOpcode(opcode_storeMLWord, cvec); decsp(); decsp(); decsp() end | BICStoreOperation { kind=LoadStoreMLByte _, address, value } => ( genByteAddress address; gencde (value, ToStack, NotEnd, loopAddr); genOpcode(opcode_storeMLByte, cvec); decsp(); decsp() ) | BICStoreOperation { kind=LoadStoreC8, address, value} => ( genByteAddress address; gencde (value, ToStack, NotEnd, loopAddr); genOpcode(opcode_storeC8, cvec); decsp(); decsp() ) | BICStoreOperation { kind=LoadStoreC16, address, value} => ( genNonByteAddress address; gencde (value, ToStack, NotEnd, loopAddr); genOpcode(opcode_storeC16, cvec); decsp(); decsp(); decsp() ) | BICStoreOperation { kind=LoadStoreC32, address, value} => ( genNonByteAddress address; gencde (value, ToStack, NotEnd, loopAddr); genOpcode(opcode_storeC32, cvec); decsp(); decsp(); decsp() ) | BICStoreOperation { kind=LoadStoreC64, address, value} => ( genNonByteAddress address; gencde (value, ToStack, NotEnd, loopAddr); genOpcode(opcode_storeC64, cvec); decsp(); decsp(); decsp() ) | BICStoreOperation { kind=LoadStoreCFloat, address, value} => ( genNonByteAddress address; gencde (value, ToStack, NotEnd, loopAddr); genOpcode(opcode_storeCFloat, cvec); decsp(); decsp(); decsp() ) | BICStoreOperation { kind=LoadStoreCDouble, address, value} => ( genNonByteAddress address; gencde (value, ToStack, NotEnd, loopAddr); genOpcode(opcode_storeCDouble, cvec); decsp(); decsp(); decsp() ) | BICStoreOperation { kind=LoadStoreUntaggedUnsigned, address, value} => ( genNonByteAddress address; gencde (value, ToStack, NotEnd, loopAddr); genOpcode(opcode_storeUntagged, cvec); decsp(); decsp(); decsp() ) | BICBlockOperation { kind=BlockOpMove{isByteMove=true}, sourceLeft, destRight, length } => ( genByteAddress sourceLeft; genByteAddress destRight; gencde (length, ToStack, NotEnd, loopAddr); genOpcode(opcode_blockMoveByte, cvec); decsp(); decsp(); decsp(); decsp() ) | BICBlockOperation { kind=BlockOpMove{isByteMove=false}, sourceLeft, destRight, length } => ( genNonByteAddress sourceLeft; genNonByteAddress destRight; gencde (length, ToStack, NotEnd, loopAddr); genOpcode(opcode_blockMoveWord, cvec); decsp(); decsp(); decsp(); decsp(); decsp(); decsp() ) | BICBlockOperation { kind=BlockOpEqualByte, sourceLeft, destRight, length } => ( genByteAddress sourceLeft; genByteAddress destRight; gencde (length, ToStack, NotEnd, loopAddr); genOpcode(opcode_blockEqualByte, cvec); decsp(); decsp(); decsp(); decsp() ) | BICBlockOperation { kind=BlockOpCompareByte, sourceLeft, destRight, length } => ( genByteAddress sourceLeft; genByteAddress destRight; gencde (length, ToStack, NotEnd, loopAddr); genOpcode(opcode_blockCompareByte, cvec); decsp(); decsp(); decsp(); decsp() ) | BICArbitrary { longCall, ... } => (* Just use the long-precision case in the interpreted version. *) ( gencde (longCall, whereto, tailKind, loopAddr) ) in (* body of gencde *) (* This ensures that there is precisely one item on the stack if whereto = ToStack and no items if whereto = NoResult. There are two points to note carefully here: (1) Negative stack adjustments are legal if we have exited. This is because matchFailFn can cut the stack back too far for its immediately enclosing expression. This is harmless because the code actually exits that expression. (2) A stack adjustment of ~1 is legal if we're generating a declaration in "ToStack" mode, because not all declarations actually generate the dummy value that we expect. This used to be handled in resetStack itself, but it's more transparent to do it here. (In addition, there was a bug in resetStack - it accumulated the stack resets, but didn't correctly accumulate these "~1" dummy value pushes.) It's all much better now. SPF 9/1/97 *) case whereto of ToStack => let val newsp = oldsp + 1; val adjustment = !realstackptr - newsp val () = if !exited orelse adjustment = 0 then () else if adjustment < ~1 then raise InternalError ("gencde: bad adjustment " ^ Int.toString adjustment) (* Hack for declarations that should push values, but don't *) else if adjustment = ~1 then pushConst (DummyValue, cvec) else resetStack (adjustment, true, cvec) in realstackptr := newsp end | NoResult => let val adjustment = !realstackptr - oldsp val () = if !exited orelse adjustment = 0 then () else if adjustment < 0 then raise InternalError ("gencde: bad adjustment " ^ Int.toString adjustment) else resetStack (adjustment, false, cvec) in realstackptr := oldsp end end (* gencde *) (* doNext is only used for mutually recursive functions where a function may not be able to fill in its closure if it does not have all the remaining declarations. *) (* TODO: This always creates the closure on the heap even when makeClosure is false. *) and genProc ({ closure=[], localCount, body, argTypes, name, ...}: bicLambdaForm, mutualDecs, doNext: unit -> unit) : unit = let (* Create a one word item for the closure. This is returned for recursive references and filled in with the address of the code when we've finished. *) val closure = makeConstantClosure() val newCode : code = codeCreate(name, parameters); (* Code-gen function. No non-local references. *) val () = codegen (body, newCode, closure, List.length argTypes, localCount, parameters); val () = pushConst(closureAsAddress closure, cvec); val () = incsp(); in if mutualDecs then doNext () else () end | genProc ({ localCount, body, name, argTypes, closure, ...}, mutualDecs, doNext) = let (* Full closure required. *) val resClosure = makeConstantClosure() val newCode = codeCreate (name, parameters) (* Code-gen function. *) val () = codegen (body, newCode, resClosure, List.length argTypes, localCount, parameters) val sizeOfClosure = List.length closure + 1; in if mutualDecs then let (* Have to make the closure now and fill it in later. *) (* This previously used genGetStore which at one time was widely used. *) val () = pushConst(toMachineWord sizeOfClosure, cvec) (* Length *) val () = pushConst(toMachineWord F_mutable, cvec) (* Flags *) val () = pushConst(toMachineWord 0, cvec) (* Initialise to zero. *) val () = genOpcode(opcode_allocWordMemory, cvec) (* Allocate the memory. *) val () = incsp () (* Put code address into closure *) val () = pushConst(codeAddressFromClosure resClosure, cvec) val () = genMoveToVec(0, cvec) val entryAddr : int = !realstackptr val () = doNext () (* Any mutually recursive functions. *) (* Push the address of the vector - If we have processed other closures the vector will no longer be on the top of the stack. *) val () = pushLocalStackValue (~ entryAddr) (* Load items for the closure. *) fun loadItems ([], _) = () | loadItems (v :: vs, addr : int) = let (* Generate an item and move it into the vector *) val () = gencde (BICExtract v, ToStack, NotEnd, NONE) val () = genMoveToVec(addr, cvec) val () = decsp () in loadItems (vs, addr + 1) end val () = loadItems (closure, 1) val () = genLock cvec (* Lock it. *) (* Remove the extra reference. *) val () = resetStack (1, false, cvec) in realstackptr := !realstackptr - 1 end else let (* Put it on the stack. *) val () = pushConst (codeAddressFromClosure resClosure, cvec) val () = incsp () val () = List.app (fn pt => gencde (BICExtract pt, ToStack, NotEnd, NONE)) closure val () = genTuple (sizeOfClosure, cvec) in realstackptr := !realstackptr - (sizeOfClosure - 1) end end and genCond (testCode, thenCode, elseCode, whereto, tailKind, loopAddr) = let val () = gencde (testCode, ToStack, NotEnd, loopAddr) val toElse = createLabel() and exitJump = createLabel() val () = putBranchInstruction(JumpFalse, toElse, cvec) val () = decsp() val () = gencde (thenCode, whereto, tailKind, loopAddr) (* Get rid of the result from the stack. If there is a result then the ``else-part'' will push it. *) val () = case whereto of ToStack => decsp () | NoResult => () val thenExited = !exited val () = if thenExited then () else putBranchInstruction (Jump, exitJump, cvec) (* start of "else part" *) val () = setLabel (toElse, cvec) val () = exited := false val () = gencde (elseCode, whereto, tailKind, loopAddr) val elseExited = !exited val () = setLabel (exitJump, cvec) in exited := (thenExited andalso elseExited) (* Only exited if both sides did. *) end (* genCond *) and genEval (eval, tailKind : tail) : unit = let val argList : backendIC list = List.map #1 (#argList eval) val argsToPass : int = List.length argList; (* Load arguments *) fun loadArgs [] = () | loadArgs (v :: vs) = let (* Push each expression onto the stack. *) val () = gencde(v, ToStack, NotEnd, NONE) in loadArgs vs end; (* Called after the args and the closure to call have been pushed onto the stack. *) fun callClosure () : unit = case tailKind of NotEnd => (* Normal call. *) genCallClosure cvec | EndOfProc => (* Tail recursive call. *) let (* Get the return address onto the top of the stack. *) val () = pushLocalStackValue 0 (* Slide the return address, closure and args over the old closure, return address and args, and reset the stack. Then jump to the closure. *) val () = genTailCall(argsToPass + 2, !realstackptr - 1 + (numOfArgs - argsToPass), cvec); (* It's "-1" not "-2", because we didn't bump the realstackptr when we pushed the return address. SPF 3/1/97 *) in exited := true end (* Have to guarantee that the expression to return the function is evaluated before the arguments. *) (* Returns true if evaluating it later is safe. *) fun safeToLeave (BICConstnt _) = true | safeToLeave (BICLambda _) = true | safeToLeave (BICExtract _) = true | safeToLeave (BICField {base, ...}) = safeToLeave base | safeToLeave (BICLoadContainer {base, ...}) = safeToLeave base | safeToLeave _ = false val () = if (case argList of [] => true | _ => safeToLeave (#function eval)) then let (* Can load the args first. *) val () = loadArgs argList in gencde (#function eval, ToStack, NotEnd, NONE) end else let (* The expression for the function is too complicated to risk leaving. It might have a side-effect and we must ensure that any side-effects it has are done before the arguments are loaded. *) val () = gencde(#function eval, ToStack, NotEnd, NONE); val () = loadArgs(argList); (* Load the function again. *) val () = genLocal(argsToPass, cvec); in incsp () end val () = callClosure () (* Call the function. *) (* Make sure we interpret when we return from the call *) val () = genEnterIntCall (cvec, argsToPass) in (* body of genEval *) realstackptr := !realstackptr - argsToPass (* Args popped by caller. *) end (* Generate the function. *) (* Assume we always want a result. There is otherwise a problem if the called routine returns a result of type void (i.e. no result) but the caller wants a result (e.g. the identity function). *) val () = gencde (pt, ToStack, EndOfProc, NONE) val () = if !exited then () else genReturn (numOfArgs, cvec); in (* body of codegen *) (* Having code-generated the body of the function, it is copied into a new data segment. *) copyCode(cvec, !maxStack, resultClosure) end (* codegen *); fun gencodeLambda({ name, body, argTypes, localCount, ...}:bicLambdaForm, parameters, closure) = let (* make the code buffer for the new function. *) val newCode : code = codeCreate (name, parameters) (* This function must have no non-local references. *) in codegen (body, newCode, closure, List.length argTypes, localCount, parameters) end local val makeEntryPoint: string -> machineWord = RunCall.rtsCallFull1 "PolyCreateEntryPointObject" fun rtsCall makeCall (entryName: string, numOfArgs, debugArgs: Universal.universal list): machineWord = let open Address val cvec = codeCreate (entryName, debugArgs) val entryPointAddr = makeEntryPoint entryName (* Each argument is at the same offset, essentially we're just shifting them *) fun genLocals 0 = () | genLocals n = (genLocal(numOfArgs +1, cvec); genLocals (n-1)) val () = genLocals numOfArgs val () = pushConst(entryPointAddr, cvec) val () = makeCall(numOfArgs, cvec) val () = genReturn (numOfArgs, cvec) val closure = makeConstantClosure() val () = copyCode(cvec, numOfArgs+1, closure) in closureAsAddress closure end in structure Foreign = struct val rtsCallFast = rtsCall genRTSCallFast and rtsCallFull = rtsCall genRTSCallFull fun rtsCallFastRealtoReal(entryName, debugArgs) = rtsCall (fn (_, c) => genRTSCallFastRealtoReal c) (entryName, 1, debugArgs) and rtsCallFastRealRealtoReal(entryName, debugArgs) = rtsCall (fn (_, c) => genRTSCallFastRealRealtoReal c) (entryName, 2, debugArgs) and rtsCallFastGeneraltoReal(entryName, debugArgs) = rtsCall (fn (_, c) => genRTSCallFastGeneraltoReal c) (entryName, 1, debugArgs) and rtsCallFastRealGeneraltoReal(entryName, debugArgs) = rtsCall (fn (_, c) => genRTSCallFastRealGeneraltoReal c) (entryName, 2, debugArgs) fun rtsCallFastFloattoFloat(entryName, debugArgs) = rtsCall (fn (_, c) => genRTSCallFastFloattoFloat c) (entryName, 1, debugArgs) and rtsCallFastFloatFloattoFloat(entryName, debugArgs) = rtsCall (fn (_, c) => genRTSCallFastFloatFloattoFloat c) (entryName, 2, debugArgs) and rtsCallFastGeneraltoFloat(entryName, debugArgs) = rtsCall (fn (_, c) => genRTSCallFastGeneraltoFloat c) (entryName, 1, debugArgs) and rtsCallFastFloatGeneraltoFloat(entryName, debugArgs) = rtsCall (fn (_, c) => genRTSCallFastFloatGeneraltoFloat c) (entryName, 2, debugArgs) end end structure Sharing = struct open BACKENDTREE.Sharing type closureRef = closureRef end end; diff --git a/mlsource/MLCompiler/CodeTree/CODETREE.ML b/mlsource/MLCompiler/CodeTree/CODETREE.ML index f97f2fad..130a6617 100644 --- a/mlsource/MLCompiler/CodeTree/CODETREE.ML +++ b/mlsource/MLCompiler/CodeTree/CODETREE.ML @@ -1,585 +1,585 @@ (* - Copyright (c) 2012,13,15-17 David C.J. Matthews + Copyright (c) 2012,13,15-17, 20 David C.J. Matthews This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) functor CODETREE ( structure DEBUG: DEBUGSIG structure PRETTY : PRETTYSIG structure BASECODETREE: BaseCodeTreeSig structure CODETREE_FUNCTIONS: CodetreeFunctionsSig structure BACKEND: sig type codetree type machineWord = Address.machineWord val codeGenerate: codetree * int * Universal.universal list -> (unit -> machineWord) * Universal.universal list structure Foreign: FOREIGNCALLSIG structure Sharing : sig type codetree = codetree end end structure OPTIMISER: sig type codetree and envSpecial and codeBinding val codetreeOptimiser: codetree * Universal.universal list * int -> { numLocals: int, general: codetree, bindings: codeBinding list, special: envSpecial } structure Sharing: sig type codetree = codetree and envSpecial = envSpecial and codeBinding = codeBinding end end sharing type PRETTY.pretty = BASECODETREE.pretty sharing BASECODETREE.Sharing = CODETREE_FUNCTIONS.Sharing = BACKEND.Sharing = OPTIMISER.Sharing ) : CODETREESIG = struct open Address; open StretchArray; open BASECODETREE; open PRETTY; open CODETREE_FUNCTIONS exception InternalError = Misc.InternalError and Interrupt = Thread.Thread.Interrupt infix 9 sub; fun mkDec (laddr, res) = Declar{value = res, addr = laddr, use=[]} fun deExtract(Extract ext) = ext | deExtract _ = raise InternalError "deExtract" datatype level = Level of { lev: int, closure: createClosure, lookup: int * int * bool -> loadForm } local (* We can have locals at the outer level. *) fun bottomLevel(addr, 0, false) = if addr < 0 then raise InternalError "load: negative" else LoadLocal addr | bottomLevel _ = (* Either the level is wrong or it's a parameter. *) raise InternalError "bottom level" in val baseLevel = Level { lev = 0, closure = makeClosure(), lookup = bottomLevel } end fun newLevel (Level{ lev, lookup, ...}) = let val closureList = makeClosure() val makeClosure = addToClosure closureList fun check n = if n < 0 then raise InternalError "load: negative" else n fun thisLevel(addr, level, isParam) = if level < 0 then raise InternalError "mkLoad: level must be non-negative" else if level > 0 then makeClosure(lookup(addr, level-1, isParam)) else (* This level *) if isParam then LoadArgument(check addr) else LoadLocal(check addr) in Level { lev = lev+1, closure = closureList, lookup = thisLevel } end fun getClosure(Level{ closure, ...}) = List.map Extract (extractClosure closure) fun mkLoad (addr, Level { lev = newLevel, lookup, ... } , Level { lev = oldLevel, ... }) = Extract(lookup(addr, newLevel - oldLevel, false)) and mkLoadParam(addr, Level { lev = newLevel, lookup, ... } , Level { lev = oldLevel, ... }) = Extract(lookup(addr, newLevel - oldLevel, true)) (* Transform a function so that free variables are converted to closure form. Returns the maximum local address used. *) fun genCode(pt, debugSwitches, numLocals) = let val printCodeTree = DEBUG.getParameter DEBUG.codetreeTag debugSwitches and compilerOut = PRETTY.getCompilerOutput debugSwitches (* val printCodeTree = true and compilerOut = PRETTY.prettyPrint(TextIO.print, 70) *) (* If required, print it first. This is the code that the front-end has produced. *) val () = if printCodeTree then compilerOut(pretty pt) else () (* This ensures that everything is printed just before it is code-generated. *) fun codeAndPrint(code, localCount) = let val () = if printCodeTree then compilerOut (BASECODETREE.pretty code) else (); in BACKEND.codeGenerate(code, localCount, debugSwitches) end (* Optimise it. *) val { numLocals = localCount, general = gen, bindings = decs, special = spec } = OPTIMISER.codetreeOptimiser(pt, debugSwitches, numLocals) (* At this stage we have a "general" value and also, possibly a "special" value. We could simply create mkEnv(decs, gen) and run preCode and genCode on that. However, we would lose the ability to insert any inline functions from this code into subsequent top-level expressions. We can't simply retain the "special" entry either because that may refer to values that have to be created once when the code is run. Such values will be referenced by "load" entries which refer to entries in the "decs". We construct a tuple which will contain the actual values after the code is run. Then if we want the value at some time in the future when we use something from the "special" entry we can extract the corresponding value from this tuple. Previously, this code always generated a tuple containing every declaration. That led to some very long compilation times because the back-end has some code which is quadratic in the number of entries on the stack. We now try to prune bindings by only generating the tuple if we have an inline function somewhere and only generating bindings we actually need. *) fun simplifySpec (EnvSpecTuple(size, env)) = let (* Get all the field entries. *) fun simpPair (gen, spec) = (gen, simplifySpec spec) val fields = List.tabulate(size, simpPair o env) in if List.all(fn (_, EnvSpecNone) => true | _ => false) fields then EnvSpecNone else EnvSpecTuple(size, fn n => List.nth(fields, n)) end | simplifySpec s = s (* None or inline function. *) in case simplifySpec spec of EnvSpecNone => let val (code, props) = codeAndPrint (mkEnv(decs, gen), localCount) in fn () => Constnt(code(), props) end | simpleSpec => let (* The bindings are marked using a three-valued mark. A binding is needed if it is referenced in any way. During the scan to find the references we need to avoid processing an entry that has already been processed but it is possible that a binding may be referenced as a general value only (e.g. from a function closure) and separately as a special value. See Test148.ML *) datatype visit = UnVisited | VisitedGeneral | VisitedSpecial local val refArray = Array.array(localCount, UnVisited) fun findDecs EnvSpecNone = () | findDecs (EnvSpecTuple(size, env)) = let val fields = List.tabulate(size, env) in List.app processGenAndSpec fields end | findDecs (EnvSpecInlineFunction({closure, ...}, env)) = let val closureItems = List.tabulate(List.length closure, env) in List.app processGenAndSpec closureItems end | findDecs (EnvSpecUnary _) = () | findDecs (EnvSpecBinary _) = () and processGenAndSpec (gen, spec) = (* The spec part needs only to be processed if this entry has not yet been visited, *) case gen of EnvGenLoad(LoadLocal addr) => let val previous = Array.sub(refArray, addr) in case (previous, spec) of (VisitedSpecial, _) => () (* Fully done *) | (VisitedGeneral, EnvSpecNone) => () (* Nothing useful *) | (_, EnvSpecNone) => (* We need this entry but we don't have any special entry to process. We could find another reference with a special entry. *) Array.update(refArray, addr, VisitedGeneral) | (_, _) => ( (* This has a special entry. Mark it and process. *) Array.update(refArray, addr, VisitedSpecial); findDecs spec ) end | EnvGenConst _ => () | _ => raise InternalError "doGeneral: not LoadLocal or Constant" val () = findDecs simpleSpec in (* Convert to an immutable data structure. This will continue to be referenced in any inline function after the code has run. *) val refVector = Array.vector refArray end val decArray = Array.array(localCount, CodeZero) fun addDec(addr, dec) = if Vector.sub(refVector, addr) <> UnVisited then Array.update(decArray, addr, dec) else () fun addDecs(Declar{addr, ...}) = addDec(addr, mkLoadLocal addr) | addDecs(RecDecs decs) = List.app(fn {addr, ...} => addDec(addr, mkLoadLocal addr)) decs | addDecs(NullBinding _) = () | addDecs(Container{addr, size, ...}) = addDec(addr, mkTupleFromContainer(addr, size)) val () = List.app addDecs decs (* Construct the tuple and add the "general" value at the start. *) val resultTuple = mkTuple(gen :: Array.foldr(op ::) nil decArray) (* Now generate the machine code and return it as a function that can be called. *) val (code, codeProps) = codeAndPrint (mkEnv (decs, resultTuple), localCount) in (* Return a function that executes the compiled code and then creates the final "global" value as the result. *) fn () => let local (* Execute the code. This will perform any side-effects the user has programmed and may raise an exception if that is required. *) val resVector = code () (* The result is a vector containing the "general" value as the first word and the evaluated bindings for any "special" entries in subsequent words. *) val decVals : address = if isShort resVector then raise InternalError "Result vector is not an address" else toAddress resVector in fun resultWordN n = loadWord (decVals, n) (* Get the general value, the zero'th entry in the vector. *) val generalVal = resultWordN 0w0 (* Get the properties for a field in the tuple. Because the result is a tuple all the properties should be contained in a tupleTag entry. *) val fieldProps = case Option.map (Universal.tagProject CodeTags.tupleTag) (List.find(Universal.tagIs CodeTags.tupleTag) codeProps) of NONE => (fn _ => []) | SOME p => (fn n => List.nth(p, n)) val generalProps = fieldProps 0 end (* Construct a new environment so that when an entry is looked up the corresponding constant is returned. *) fun newEnviron (oldEnv) args = let val (oldGeneral, oldSpecial) = oldEnv args val genPair = case oldGeneral of EnvGenLoad(LoadLocal addr) => ( (* For the moment retain this check. It's better to have an assertion failure than a segfault. *) Vector.sub(refVector, addr) <> UnVisited orelse raise InternalError "Reference to non-existent binding"; (resultWordN(Word.fromInt addr+0w1), fieldProps(addr+1)) ) | EnvGenConst c => c | _ => raise InternalError "codetree newEnviron: Not Extract or Constnt" val specVal = mapSpec oldSpecial in (EnvGenConst genPair, specVal) end and mapSpec EnvSpecNone = EnvSpecNone | mapSpec (EnvSpecTuple(size, env)) = EnvSpecTuple(size, newEnviron env) | mapSpec (EnvSpecInlineFunction(spec, env)) = EnvSpecInlineFunction(spec, (newEnviron env)) | mapSpec (EnvSpecUnary _) = EnvSpecNone | mapSpec (EnvSpecBinary _) = EnvSpecNone in (* and return the whole lot as a global value. *) Constnt(generalVal, setInline(mapSpec simpleSpec) generalProps) end end end (* genCode *) (* Constructor functions for the front-end of the compiler. *) local fun mkSimpleFunction inlineType (lval, args, name, closure, numLocals) = { body = lval, isInline = inlineType, name = if name = "" then "" else name, closure = map deExtract closure, argTypes = List.tabulate(args, fn _ => (GeneralType, [])), resultType = GeneralType, localCount = numLocals, recUse = [] } in val mkProc = Lambda o mkSimpleFunction NonInline (* Normal function *) and mkInlproc = Lambda o mkSimpleFunction Inline (* Explicitly inlined by the front-end *) (* Unless Compiler.inlineFunctor is false functors are treated as macros and expanded when they are applied. Unlike core-language functions they are not first-class values so if they are inline the "value" returned in the initial binding can just be zero except if there is something in the closure. Almost always the closure will be empty since free variables will come from previous topdecs and will be constants, The exception is if a structure and a functor using the structure appear in the same topdec (no semicolon between them). In that case we can't leave it. We would have to update the closure even if we leave the body untouched but we could have closure entries that are constants. e.g. structure S = struct val x = 1 end functor F() = struct open S end *) fun mkMacroProc (args as (_, _, _, [], _)) = Constnt(toMachineWord 0, setInline ( EnvSpecInlineFunction(mkSimpleFunction Inline args, fn _ => raise InternalError "mkMacroProc: closure")) []) | mkMacroProc args = Lambda(mkSimpleFunction Inline args) end local fun mkFunWithTypes inlineType { body, argTypes=argsAndTypes, resultType, name, closure, numLocals } = Lambda { body = body, isInline = inlineType, name = if name = "" then "" else name, closure = map deExtract closure, argTypes = map (fn t => (t, [])) argsAndTypes, resultType = resultType, localCount = numLocals, recUse = [] } in val mkFunction = mkFunWithTypes NonInline and mkInlineFunction = mkFunWithTypes Inline end fun mkEval (ct, clist) = Eval { function = ct, argList = List.map(fn c => (c, GeneralType)) clist, resultType=GeneralType } fun mkCall(func, argsAndTypes, resultType) = Eval { function = func, argList = argsAndTypes, resultType=resultType } (* Basic built-in operations. *) fun mkUnary (oper, arg1) = Unary { oper = oper, arg1 = arg1 } and mkBinary (oper, arg1, arg2) = Binary { oper = oper, arg1 = arg1, arg2 = arg2 } val getCurrentThreadId = GetThreadId and getCurrentThreadIdFn = mkInlproc(GetThreadId, 1 (* Ignores argument *), "GetThreadId()", [], 0) fun mkAllocateWordMemory (numWords, flags, initial) = AllocateWordMemory { numWords = numWords, flags = flags, initial = initial } val mkAllocateWordMemoryFn = mkInlproc( mkAllocateWordMemory(mkInd(0, mkLoadArgument 0), mkInd(1, mkLoadArgument 0), mkInd(2, mkLoadArgument 0)), 1, "AllocateWordMemory()", [], 0) (* Builtins wrapped as functions. N.B. These all take a single argument which may be a tuple. *) fun mkUnaryFn oper = mkInlproc(mkUnary(oper, mkLoadArgument 0), 1, BuiltIns.unaryRepr oper ^ "()", [], 0) and mkBinaryFn oper = mkInlproc(mkBinary(oper, mkInd(0, mkLoadArgument 0), mkInd(1, mkLoadArgument 0)), 1, BuiltIns.binaryRepr oper ^ "()", [], 0) local open BuiltIns (* Word equality. The value of isSigned doesn't matter. *) val eqWord = WordComparison{test=TestEqual, isSigned=false} in fun mkNot arg = Unary{oper=NotBoolean, arg1=arg} and mkIsShort arg = Unary{oper=IsTaggedValue, arg1=arg} - and mkEqualWord (arg1, arg2) = + and mkEqualTaggedWord (arg1, arg2) = Binary{oper=eqWord, arg1=arg1, arg2=arg2} - val equalWordFn = (* This takes two words, not a tuple. *) + and mkEqualPointerOrWord (arg1, arg2) = + Binary{oper=PointerEq, arg1=arg1, arg2=arg2} + val equalTaggedWordFn = (* This takes two words, not a tuple. *) mkInlproc(mkBinary(eqWord, mkLoadArgument 0, mkLoadArgument 1), 2, "EqualWord()", [], 0) + and equalPointerOrWordFn = (* This takes two words, not a tuple. *) + mkInlproc(mkBinary(PointerEq, mkLoadArgument 0, mkLoadArgument 1), 2, "EqualWord()", [], 0) end - - (* Equality for arbitrary precision if at least one of the arguments is known to be short. *) - fun mkEqualArbShort (arg1, arg2) = - Arbitrary { oper=ArbCompare BuiltIns.TestEqual, shortCond=Constnt(toMachineWord 1, []), arg1=arg1, arg2=arg2, longCall=CodeZero} fun mkLoadOperation(oper, base, index) = LoadOperation{kind=oper, address={base=base, index=SOME index, offset=0w0}} fun mkLoadOperationFn oper = mkInlproc(mkLoadOperation(oper, mkInd(0, mkLoadArgument 0), mkInd(1, mkLoadArgument 0)), 1, "loadOperation()", [], 0) fun mkStoreOperation(oper, base, index, value) = StoreOperation{kind=oper, address={base=base, index=SOME index, offset=0w0}, value=value} fun mkStoreOperationFn oper = mkInlproc(mkStoreOperation(oper, mkInd(0, mkLoadArgument 0), mkInd(1, mkLoadArgument 0), mkInd(2, mkLoadArgument 0)), 1, "storeOperation()", [], 0) fun mkBlockOperation {kind, leftBase, leftIndex, rightBase, rightIndex, length } = BlockOperation { kind = kind, sourceLeft={base=leftBase, index=SOME leftIndex, offset=0w0}, destRight={base=rightBase, index=SOME rightIndex, offset=0w0}, length=length} (* Construct a function that takes five arguments. The order is left-base, right-base, left-index, right-index, length. *) fun mkBlockOperationFn kind = mkInlproc( mkBlockOperation{kind=kind, leftBase=mkInd(0, mkLoadArgument 0), rightBase=mkInd(1, mkLoadArgument 0), leftIndex=mkInd(2, mkLoadArgument 0), rightIndex=mkInd(3, mkLoadArgument 0), length=mkInd(4, mkLoadArgument 0)}, 1, "blockOperation()", [], 0) fun identityFunction (name : string) : codetree = mkInlproc (mkLoadArgument 0, 1, name, [], 0) (* Returns its argument. *); (* Test a tag value. *) fun mkTagTest(test: codetree, tagValue: word, maxTag: word) = TagTest {test=test, tag=tagValue, maxTag=maxTag } fun mkHandle (exp, handler, exId) = Handle {exp = exp, handler = handler, exPacketAddr = exId} fun mkStr (strbuff:string) = Constnt (toMachineWord strbuff, []) (* If we have multiple references to a piece of code we may have to save it in a temporary and then use it from there. If the code has side-effects we certainly must do that to ensure that the side-effects are done exactly once and in the correct order, however if the code is just a constant or a load we can reduce the amount of code we generate by simply returning the original code. *) fun multipleUses (code as Constnt _, _, _) = {load = (fn _ => code), dec = []} (* | multipleUses (code as Extract(LoadLegacy{addr, level=loadLevel, ...}), _, level) = let (* May have to adjust the level. *) fun loadFn lev = if lev = level then code else mkLoad (addr, loadLevel + lev, level)) in {load = loadFn, dec = []} end | multipleUses (code as Extract(LoadLocal addr), _, level) = let (* May have to adjust the level. *) fun loadFn lev = if lev = level then code else mkLoad (addr, lev - level) in {load = loadFn, dec = []} end | multipleUses (code as Extract(LoadArgument _), _, level) = let (* May have to adjust the level. *) fun loadFn lev = if lev = level then code else raise InternalError "multipleUses: different level" (*else mkLoad (addr, lev - level)*) in {load = loadFn, dec = []} end | multipleUses (Extract _, _, _) = raise InternalError "multipleUses: TODO" *) | multipleUses (code, nextAddress, level) = let val addr = nextAddress(); fun loadFn lev = mkLoad (addr, lev, level); in {load = loadFn, dec = [mkDec (addr, code)]} end (* multipleUses *); fun mkMutualDecs [] = raise InternalError "mkMutualDecs: empty declaration list" | mkMutualDecs l = let fun convertDec(a, Lambda lam) = {lambda = lam, addr = a, use=[]} | convertDec _ = raise InternalError "mkMutualDecs: Recursive declaration is not a function" in RecDecs(List.map convertDec l) end val mkNullDec = NullBinding fun mkContainer(addr, size, setter) = Container{addr=addr, size=size, use=[], setter=setter} val mkIf = Cond and mkRaise = Raise fun mkConst v = Constnt(v, []) (* For the moment limit these to general arguments. *) fun mkLoop args = Loop (List.map(fn c => (c, GeneralType)) args) and mkBeginLoop(exp, args) = BeginLoop{loop=exp, arguments=List.map(fn(i, v) => ({value=v, addr=i, use=[]}, GeneralType)) args} fun mkWhile(b, e) = (* Generated as if b then (e; ) else (). *) mkBeginLoop(mkIf(b, mkEnv([NullBinding e], mkLoop[]), CodeZero), []) (* We previously had conditional-or and conditional-and as separate instructions. I've taken them out since they can be implemented just as efficiently as a normal conditional. In addition they were interfering with the optimisation where the second expression contained the last reference to something. We needed to add a "kill entry" to the other branch but there wasn't another branch to add it to. DCJM 7/12/00. *) fun mkCor(xp1, xp2) = mkIf(xp1, CodeTrue, xp2); fun mkCand(xp1, xp2) = mkIf(xp1, xp2, CodeZero); val mkSetContainer = fn (container, tuple, size) => mkSetContainer(container, tuple, BoolVector.tabulate(size, fn _ => true)) (* An arbitrary precision operation takes a tuple consisting of a pair of arguments and a function. The code that is constructed checks both arguments to see if they are short. If they are not or the short precision operation overflows the code to call the function is executed. *) fun mkArbitraryFn oper = mkInlproc( Arbitrary{oper=oper, shortCond=mkCand(mkIsShort(mkInd(0, mkLoadArgument 0)), mkIsShort(mkInd(1, mkLoadArgument 0))), arg1=mkInd(0, mkLoadArgument 0), arg2=mkInd(1, mkLoadArgument 0), longCall= mkEval(mkInd(2, mkLoadArgument 0), [mkTuple[mkInd(0, mkLoadArgument 0), mkInd(1, mkLoadArgument 0)]])}, 1, "Arbitrary" ^ (case oper of ArbCompare test => BuiltIns.testRepr test | ArbArith arith => BuiltIns.arithRepr arith) ^ "()", [], 0) structure Foreign = BACKEND.Foreign structure Sharing = struct type machineWord = machineWord type codetree = codetree type pretty = pretty type argumentType=argumentType type codeBinding = codeBinding type level = level end end (* CODETREE functor body *); diff --git a/mlsource/MLCompiler/CodeTree/CODETREE_FUNCTIONS.sml b/mlsource/MLCompiler/CodeTree/CODETREE_FUNCTIONS.sml index 7ecb6795..b8e8d1ad 100644 --- a/mlsource/MLCompiler/CodeTree/CODETREE_FUNCTIONS.sml +++ b/mlsource/MLCompiler/CodeTree/CODETREE_FUNCTIONS.sml @@ -1,492 +1,493 @@ (* Copyright (c) 2012,13,16,18-20 David C.J. Matthews This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) (* Miscellaneous construction and operation functions on the code-tree. *) functor CODETREE_FUNCTIONS( structure BASECODETREE: BaseCodeTreeSig structure STRONGLY: sig val stronglyConnectedComponents: {nodeAddress: 'a -> int, arcs: 'a -> int list } -> 'a list -> 'a list list end ) : CodetreeFunctionsSig = struct open BASECODETREE open STRONGLY open Address exception InternalError = Misc.InternalError fun mkEnv([], exp) = exp | mkEnv(decs, exp) = Newenv(decs, exp) val word0 = toMachineWord 0 and word1 = toMachineWord 1 val False = word0 and True = word1 val F_mutable_words : Word8.word = Word8.orb (F_words, F_mutable) val CodeFalse = Constnt(False, []) and CodeTrue = Constnt(True, []) and CodeZero = Constnt(word0, []) (* Properties of code. This indicates the extent to which the code has side-effects (i.e. where even if the result is unused the code still needs to be produced) or is applicative (i.e. where its value depends only arguments and can safely be reordered). *) (* The RTS has a table of properties for RTS functions. The 103 call returns these Or-ed into the register mask. *) val PROPWORD_NORAISE = 0wx40000000 and PROPWORD_NOUPDATE = 0wx20000000 and PROPWORD_NODEREF = 0wx10000000 (* Since RTS calls are being eliminated leave residual versions of these. *) fun earlyRtsCall _ = false and sideEffectFreeRTSCall _ = false local infix orb andb val op orb = Word.orb and op andb = Word.andb val noSideEffect = PROPWORD_NORAISE orb PROPWORD_NOUPDATE val applicative = noSideEffect orb PROPWORD_NODEREF in fun codeProps (Lambda _) = applicative | codeProps (Constnt _) = applicative | codeProps (Extract _) = applicative | codeProps (TagTest{ test, ... }) = codeProps test | codeProps (Cond(i, t, e)) = codeProps i andb codeProps t andb codeProps e | codeProps (Newenv(decs, exp)) = List.foldl (fn (d, r) => bindingProps d andb r) (codeProps exp) decs | codeProps (Handle { exp, handler, ... }) = (* A handler processes all the exceptions in the body *) (codeProps exp orb PROPWORD_NORAISE) andb codeProps handler | codeProps (Tuple { fields, ...}) = testList fields | codeProps (Indirect{base, ...}) = codeProps base (* A built-in function may be side-effect free. This can occur if we have, for example, "if exp1 orelse exp2" where exp2 can be reduced to "true", typically because it's inside an inline function and some of the arguments to the function are constants. This then gets converted to (exp1; true) and we can eliminate exp1 if it is simply a comparison. *) | codeProps GetThreadId = Word.orb(PROPWORD_NOUPDATE, PROPWORD_NORAISE) | codeProps (Unary{oper, arg1}) = let open BuiltIns val operProps = case oper of NotBoolean => applicative | IsTaggedValue => applicative | MemoryCellLength => applicative (* MemoryCellFlags could return a different result if a mutable cell was locked. *) | MemoryCellFlags => applicative | ClearMutableFlag => Word.orb(PROPWORD_NODEREF, PROPWORD_NORAISE) | AtomicIncrement => PROPWORD_NORAISE | AtomicDecrement => PROPWORD_NORAISE | AtomicReset => Word.orb(PROPWORD_NODEREF, PROPWORD_NORAISE) | LongWordToTagged => applicative | SignedToLongWord => applicative | UnsignedToLongWord => applicative | RealAbs _ => applicative (* Does not depend on rounding setting. *) | RealNeg _ => applicative (* Does not depend on rounding setting. *) (* If we float a 64-bit int to a 64-bit floating point value we may lose precision so this depends on the current rounding mode. *) | RealFixedInt _ => Word.orb(PROPWORD_NOUPDATE, PROPWORD_NORAISE) | FloatToDouble => applicative (* The rounding mode is set explicitly. *) | DoubleToFloat _ => applicative (* May raise the overflow exception *) | RealToInt _ => PROPWORD_NOUPDATE orb PROPWORD_NODEREF | TouchAddress => PROPWORD_NORAISE (* Treat as updating a notional reference count. *) in operProps andb codeProps arg1 end | codeProps (Binary{oper, arg1, arg2}) = let open BuiltIns val mayRaise = PROPWORD_NOUPDATE orb PROPWORD_NODEREF val operProps = case oper of WordComparison _ => applicative | FixedPrecisionArith _ => mayRaise | WordArith _ => applicative (* Quot and Rem don't raise exceptions - zero checking is done before. *) | WordLogical _ => applicative | WordShift _ => applicative | AllocateByteMemory => Word.orb(PROPWORD_NOUPDATE, PROPWORD_NORAISE) (* Allocation returns a different value on each call. *) | LargeWordComparison _ => applicative | LargeWordArith _ => applicative (* Quot and Rem don't raise exceptions - zero checking is done before. *) | LargeWordLogical _ => applicative | LargeWordShift _ => applicative | RealComparison _ => applicative (* Real arithmetic operations depend on the current rounding setting. *) | RealArith _ => Word.orb(PROPWORD_NOUPDATE, PROPWORD_NORAISE) + | PointerEq => applicative in operProps andb codeProps arg1 andb codeProps arg2 end | codeProps (Arbitrary{shortCond, arg1, arg2, longCall, ...}) = (* Arbitrary precision operations are applicative but the longCall is a function call. It should never have a side-effect so it might be better to remove it. *) codeProps shortCond andb codeProps arg1 andb codeProps arg2 andb codeProps longCall | codeProps (AllocateWordMemory {numWords, flags, initial}) = let val operProps = Word.orb(PROPWORD_NOUPDATE, PROPWORD_NORAISE) in operProps andb codeProps numWords andb codeProps flags andb codeProps initial end | codeProps (Eval _) = 0w0 | codeProps(Raise exp) = codeProps exp andb (Word.notb PROPWORD_NORAISE) (* Treat these as unsafe at least for the moment. *) | codeProps(BeginLoop _) = 0w0 | codeProps(Loop _) = 0w0 | codeProps (SetContainer _) = 0w0 | codeProps (LoadOperation {address, kind}) = let val operProps = case kind of LoadStoreMLWord {isImmutable=true} => applicative | LoadStoreMLByte {isImmutable=true} => applicative | _ => Word.orb(PROPWORD_NOUPDATE, PROPWORD_NORAISE) in operProps andb addressProps address end | codeProps (StoreOperation {address, value, ...}) = Word.orb(PROPWORD_NODEREF, PROPWORD_NORAISE) andb addressProps address andb codeProps value | codeProps (BlockOperation {kind, sourceLeft, destRight, length}) = let val operProps = case kind of BlockOpMove _ => PROPWORD_NORAISE | BlockOpEqualByte => applicative | BlockOpCompareByte => applicative in operProps andb addressProps sourceLeft andb addressProps destRight andb codeProps length end and testList t = List.foldl(fn (c, r) => codeProps c andb r) applicative t and bindingProps(Declar{value, ...}) = codeProps value | bindingProps(RecDecs _) = applicative (* These should all be lambdas *) | bindingProps(NullBinding c) = codeProps c | bindingProps(Container{setter, ...}) = codeProps setter and addressProps{base, index=NONE, ...} = codeProps base | addressProps{base, index=SOME index, ...} = codeProps base andb codeProps index (* sideEffectFree - does not raise an exception or make an assignment. *) fun sideEffectFree c = (codeProps c andb noSideEffect) = noSideEffect (* reorderable - does not raise an exception or access a reference. *) and reorderable c = codeProps c = applicative end (* Return the inline property if it is set. *) fun findInline [] = EnvSpecNone | findInline (h::t) = if Universal.tagIs CodeTags.inlineCodeTag h then Universal.tagProject CodeTags.inlineCodeTag h else findInline t (* Makes a constant value from an expression which is known to be constant but may involve inline functions, tuples etc. *) fun makeConstVal (cVal:codetree) = let fun makeVal (c as Constnt _) = c (* should just be a tuple *) (* Get a vector, copy the entries into it and return it as a constant. *) | makeVal (Tuple {fields= [], ...}) = CodeZero (* should have been optimised already! *) | makeVal (Tuple {fields, ...}) = let val tupleSize = List.length fields val vec : address = allocWordData(Word.fromInt tupleSize, F_mutable_words, word0) val fieldCode = map makeVal fields fun copyToVec ([], _) = [] | copyToVec (Constnt(w, prop) :: t, locn) = ( assignWord (vec, locn, w); prop :: copyToVec (t, locn + 0w1) ) | copyToVec _ = raise InternalError "not constant" val props = copyToVec(fieldCode, 0w0) (* If any of the constants have properties create a tuple property for the result. *) val tupleProps = if List.all null props then [] else let (* We also need to construct an EnvSpecTuple property because findInline does not look at tuple properties. *) val inlineProps = map findInline props val inlineProp = if List.all (fn EnvSpecNone => true | _ => false) inlineProps then [] else let fun tupleEntry n = (EnvGenConst(loadWord(vec, Word.fromInt n), List.nth(props, n)), List.nth(inlineProps, n)) in [Universal.tagInject CodeTags.inlineCodeTag (EnvSpecTuple(tupleSize, tupleEntry))] end in Universal.tagInject CodeTags.tupleTag props :: inlineProp end in lock vec; Constnt(toMachineWord vec, tupleProps) end | makeVal _ = raise InternalError "makeVal - not constant or tuple" in makeVal cVal end local fun allConsts [] = true | allConsts (Constnt _ :: t) = allConsts t | allConsts _ = false fun mkRecord isVar xp = let val tuple = Tuple{fields = xp, isVariant = isVar } in if allConsts xp then (* Make it now. *) makeConstVal tuple else tuple end; in val mkTuple = mkRecord false and mkDatatype = mkRecord true end (* Set the inline property. If the property is already present it is replaced. If the property we are setting is EnvSpecNone no property is set. *) fun setInline p (h::t) = if Universal.tagIs CodeTags.inlineCodeTag h then setInline p t else h :: setInline p t | setInline EnvSpecNone [] = [] | setInline p [] = [Universal.tagInject CodeTags.inlineCodeTag p] (* These are very frequently used and it might be worth making special bindings for values such as 0, 1, 2, 3 etc to reduce garbage. *) fun checkNonZero n = if n < 0 then raise InternalError "mkLoadxx: argument negative" else n val mkLoadLocal = Extract o LoadLocal o checkNonZero and mkLoadArgument = Extract o LoadArgument o checkNonZero and mkLoadClosure = Extract o LoadClosure o checkNonZero (* Set the container to the fields of the record. Try to push this down as far as possible. *) fun mkSetContainer(container, Cond(ifpt, thenpt, elsept), filter) = Cond(ifpt, mkSetContainer(container, thenpt, filter), mkSetContainer(container, elsept, filter)) | mkSetContainer(container, Newenv(decs, exp), filter) = Newenv(decs, mkSetContainer(container, exp, filter)) | mkSetContainer(_, r as Raise _, _) = r (* We may well have the situation where one branch of an "if" raises an exception. We can simply raise the exception on that branch. *) | mkSetContainer(container, Handle {exp, handler, exPacketAddr}, filter) = Handle{exp=mkSetContainer(container, exp, filter), handler=mkSetContainer(container, handler, filter), exPacketAddr = exPacketAddr} | mkSetContainer(container, tuple, filter) = SetContainer{container = container, tuple = tuple, filter = filter } local val except: exn = InternalError "Invalid load encountered in compiler" (* Exception value to use for invalid cases. We put this in the code but it should never actually be executed. *) val raiseError = Raise (Constnt (toMachineWord except, [])) in (* Look for an entry in a tuple. Used in both the optimiser and in mkInd. *) fun findEntryInBlock (Tuple { fields, isVariant, ...}, offset, isVar) = ( isVariant = isVar orelse raise InternalError "findEntryInBlock: tuple/datatype mismatch"; if offset < List.length fields then List.nth(fields, offset) (* This can arise if we're processing a branch of a case discriminating on a datatype which won't actually match at run-time. e.g. Tests/Succeed/Test030. *) else if isVar then raiseError else raise InternalError "findEntryInBlock: invalid address" ) | findEntryInBlock (Constnt (b, props), offset, isVar) = let (* Find the tuple property if it is present and extract the field props. *) val fieldProps = case List.find(Universal.tagIs CodeTags.tupleTag) props of NONE => [] | SOME p => List.nth(Universal.tagProject CodeTags.tupleTag p, offset) in case findInline props of EnvSpecTuple(_, env) => (* Do the selection now. This is especially useful if we have a global structure *) (* At the moment at least we assume that we can get all the properties from the tuple selection. *) ( case env offset of (EnvGenConst(w, p), inl) => Constnt(w, setInline inl p) (* The general value from selecting a field from a constant tuple must be a constant. *) | _ => raise InternalError "findEntryInBlock: not constant" ) | _ => (* The ML compiler may generate loads from invalid addresses as a result of a val binding to a constant which has the wrong shape. e.g. val a :: b = nil It will always result in a Bind exception being generated before the invalid load, but we have to be careful that the optimiser does not fall over. *) if isShort b orelse not (Address.isWords (toAddress b)) orelse Address.length (toAddress b) <= Word.fromInt offset then if isVar then raiseError else raise InternalError "findEntryInBlock: invalid address" else Constnt (loadWord (toAddress b, Word.fromInt offset), fieldProps) end | findEntryInBlock(base, offset, isVar) = Indirect {base = base, offset = offset, indKind = if isVar then IndVariant else IndTuple} (* anything else *) end (* Exported indirect load operation i.e. load a field from a tuple. We can't use findEntryInBlock in every case since that discards unused entries in a tuple and at this point we haven't checked that the unused entries don't have side-effects/raise exceptions e.g. #1 (1, raise Fail "bad") *) local fun mkIndirect isVar (addr, base as Constnt _) = findEntryInBlock(base, addr, isVar) | mkIndirect isVar (addr, base) = Indirect {base = base, offset = addr, indKind = if isVar then IndVariant else IndTuple} in val mkInd = mkIndirect false and mkVarField = mkIndirect true end fun mkIndContainer(addr, base) = Indirect{offset=addr, base=base, indKind=IndContainer} (* Create a tuple from a container. *) fun mkTupleFromContainer(addr, size) = Tuple{fields = List.tabulate(size, fn n => mkIndContainer(n, mkLoadLocal addr)), isVariant = false} (* Get the value from the code. *) fun evalue (Constnt(c, _)) = SOME c | evalue _ = NONE (* This is really to simplify the change from mkEnv taking a codetree list to taking a codeBinding list * code. This extracts the last entry which must be a NullBinding and packages the declarations with it. *) fun decSequenceWithFinalExp decs = let fun splitLast _ [] = raise InternalError "decSequenceWithFinalExp: empty" | splitLast decs [NullBinding exp] = (List.rev decs, exp) | splitLast _ [_] = raise InternalError "decSequenceWithFinalExp: last is not a NullDec" | splitLast decs (hd::tl) = splitLast (hd:: decs) tl in mkEnv(splitLast [] decs) end local type node = { addr: int, lambda: lambdaForm, use: codeUse list } fun nodeAddress({addr, ...}: node) = addr and arcs({lambda={closure, ...}, ...}: node) = List.foldl(fn (LoadLocal addr, l) => addr :: l | (_, l) => l) [] closure in val stronglyConnected = stronglyConnectedComponents{nodeAddress=nodeAddress, arcs=arcs} end (* In general any mutually recursive declaration can refer to any other. It's better to partition the recursive declarations into strongly connected components i.e. those that actually refer to each other. *) fun partitionMutualBindings(RecDecs rlist) = let val processed = stronglyConnected rlist (* Convert the result. Note that stronglyConnectedComponents returns the dependencies in the reverse order i.e. if X depends on Y but not the other way round then X will appear before Y in the list. We need to reverse it so that X goes after Y. *) fun rebuild ([{lambda, addr, use}], tl) = Declar{addr=addr, use=use, value=Lambda lambda} :: tl | rebuild (multiple, tl) = RecDecs multiple :: tl in List.foldl rebuild [] processed end (* This is only intended for RecDecs but it's simpler to handle all bindings. *) | partitionMutualBindings other = [other] (* Functions to help in building a closure. *) datatype createClosure = Closure of (loadForm * int) list ref fun makeClosure() = Closure(ref []) (* Function to build a closure. Items are added to the closure if they are not already there. *) fun addToClosure (Closure closureList) (ext: loadForm): loadForm = case (List.find (fn (l, _) => l = ext) (!closureList), ! closureList) of (SOME(_, n), _) => (* Already there *) LoadClosure n | (NONE, []) => (* Not there - first *) (closureList := [(ext, 0)]; LoadClosure 0) | (NONE, cl as (_, n) :: _) => (closureList := (ext, n+1) :: cl; LoadClosure(n+1)) fun extractClosure(Closure (ref closureList)) = List.foldl (fn ((ext, _), l) => ext :: l) [] closureList structure Sharing = struct type codetree = codetree and codeBinding = codeBinding and loadForm = loadForm and createClosure = createClosure and envSpecial = envSpecial end end; diff --git a/mlsource/MLCompiler/CodeTree/CODETREE_SIMPLIFIER.sml b/mlsource/MLCompiler/CodeTree/CODETREE_SIMPLIFIER.sml index 389069e8..d8ab208e 100644 --- a/mlsource/MLCompiler/CodeTree/CODETREE_SIMPLIFIER.sml +++ b/mlsource/MLCompiler/CodeTree/CODETREE_SIMPLIFIER.sml @@ -1,1739 +1,1745 @@ (* Copyright (c) 2013, 2016-17, 2020 David C.J. Matthews This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) (* This is a cut-down version of the optimiser which simplifies the code but does not apply any heuristics. It follows chained bindings, in particular through tuples, folds constants expressions involving built-in functions, expands inline functions that have previously been marked as inlineable. It does not detect small functions that can be inlined nor does it code-generate functions without free variables. *) functor CODETREE_SIMPLIFIER( structure BASECODETREE: BaseCodeTreeSig structure CODETREE_FUNCTIONS: CodetreeFunctionsSig structure REMOVE_REDUNDANT: sig type codetree type loadForm type codeUse val cleanProc : (codetree * codeUse list * (int -> loadForm) * int) -> codetree structure Sharing: sig type codetree = codetree and loadForm = loadForm and codeUse = codeUse end end sharing BASECODETREE.Sharing = CODETREE_FUNCTIONS.Sharing = REMOVE_REDUNDANT.Sharing ) : sig type codetree and codeBinding and envSpecial val simplifier: codetree * int -> (codetree * codeBinding list * envSpecial) * int * bool val specialToGeneral: codetree * codeBinding list * envSpecial -> codetree structure Sharing: sig type codetree = codetree and codeBinding = codeBinding and envSpecial = envSpecial end end = struct open BASECODETREE open Address open CODETREE_FUNCTIONS open BuiltIns exception InternalError = Misc.InternalError exception RaisedException (* The bindings are held internally as a reversed list. This is really only a check that the reversed and forward lists aren't confused. *) datatype revlist = RevList of codeBinding list type simpContext = { lookupAddr: loadForm -> envGeneral * envSpecial, enterAddr: int * (envGeneral * envSpecial) -> unit, nextAddress: unit -> int, reprocess: bool ref } fun envGeneralToCodetree(EnvGenLoad ext) = Extract ext | envGeneralToCodetree(EnvGenConst w) = Constnt w fun mkDec (laddr, res) = Declar{value = res, addr = laddr, use=[]} fun mkEnv([], exp) = exp | mkEnv(decs, exp as Extract(LoadLocal loadAddr)) = ( (* A common case is where we have a binding as the last item and then a load of that binding. Reduce this so other optimisations are possible. This is still something of a special case that could/should be generalised. *) case List.last decs of Declar{addr=decAddr, value, ... } => if loadAddr = decAddr then mkEnv(List.take(decs, List.length decs - 1), value) else Newenv(decs, exp) | _ => Newenv(decs, exp) ) | mkEnv(decs, exp) = Newenv(decs, exp) fun isConstnt(Constnt _) = true | isConstnt _ = false (* Wrap up the general, bindings and special value as a codetree node. The special entry is discarded except for Constnt entries which are converted to ConstntWithInline. That allows any inlineable code to be carried forward to later passes. *) fun specialToGeneral(g, RevList(b as _ :: _), s) = mkEnv(List.rev b, specialToGeneral(g, RevList [], s)) | specialToGeneral(Constnt(w, p), RevList [], s) = Constnt(w, setInline s p) | specialToGeneral(g, RevList [], _) = g (* Convert a constant to a fixed value. Used in some constant folding. *) val toFix: machineWord -> FixedInt.int = FixedInt.fromInt o Word.toIntX o toShort local val ffiSizeFloat: unit -> word = RunCall.rtsCallFast1 "PolySizeFloat" and ffiSizeDouble: unit -> word = RunCall.rtsCallFast1 "PolySizeDouble" in (* If we have a constant index value we convert that into a byte offset. We need to know the size of the item on this platform. We have to make this check when we actually compile the code because the interpreted version will generally be run on a platform different from the one the pre-built compiler was compiled on. The ML word length will be the same because we have separate pre-built compilers for 32 and 64-bit. *) fun getMultiplier (LoadStoreMLWord _) = RunCall.bytesPerWord | getMultiplier (LoadStoreMLByte _) = 0w1 | getMultiplier LoadStoreC8 = 0w1 | getMultiplier LoadStoreC16 = 0w2 | getMultiplier LoadStoreC32 = 0w4 | getMultiplier LoadStoreC64 = 0w8 | getMultiplier LoadStoreCFloat = ffiSizeFloat() | getMultiplier LoadStoreCDouble = ffiSizeDouble() | getMultiplier LoadStoreUntaggedUnsigned = RunCall.bytesPerWord end fun simplify(c, s) = mapCodetree (simpGeneral s) c (* Process the codetree to return a codetree node. This is used when we don't want the special case. *) and simpGeneral { lookupAddr, ...} (Extract ext) = let val (gen, spec) = lookupAddr ext in SOME(specialToGeneral(envGeneralToCodetree gen, RevList [], spec)) end | simpGeneral context (Newenv envArgs) = SOME(specialToGeneral(simpNewenv(envArgs, context, RevList []))) | simpGeneral context (Lambda lambda) = SOME(Lambda(#1(simpLambda(lambda, context, NONE, NONE)))) | simpGeneral context (Eval {function, argList, resultType}) = SOME(specialToGeneral(simpFunctionCall(function, argList, resultType, context, RevList[]))) (* BuiltIn0 functions can't be processed specially. *) | simpGeneral context (Unary{oper, arg1}) = SOME(specialToGeneral(simpUnary(oper, arg1, context, RevList []))) | simpGeneral context (Binary{oper, arg1, arg2}) = SOME(specialToGeneral(simpBinary(oper, arg1, arg2, context, RevList []))) | simpGeneral context (Arbitrary{oper=ArbCompare test, shortCond, arg1, arg2, longCall}) = SOME(specialToGeneral(simpArbitraryCompare(test, shortCond, arg1, arg2, longCall, context, RevList []))) | simpGeneral context (Arbitrary{oper=ArbArith arith, shortCond, arg1, arg2, longCall}) = SOME(specialToGeneral(simpArbitraryArith(arith, shortCond, arg1, arg2, longCall, context, RevList []))) | simpGeneral context (AllocateWordMemory {numWords, flags, initial}) = SOME(specialToGeneral(simpAllocateWordMemory(numWords, flags, initial, context, RevList []))) | simpGeneral context (Cond(condTest, condThen, condElse)) = SOME(specialToGeneral(simpIfThenElse(condTest, condThen, condElse, context, RevList []))) | simpGeneral context (Tuple { fields, isVariant }) = SOME(specialToGeneral(simpTuple(fields, isVariant, context, RevList []))) | simpGeneral context (Indirect{ base, offset, indKind }) = SOME(specialToGeneral(simpFieldSelect(base, offset, indKind, context, RevList []))) | simpGeneral context (SetContainer{container, tuple, filter}) = let val optCont = simplify(container, context) val (cGen, cDecs, cSpec) = simpSpecial(tuple, context, RevList []) in case cSpec of (* If the tuple is a local binding it is simpler to pick it up from the "special" entry. *) EnvSpecTuple(size, recEnv) => let val fields = List.tabulate(size, envGeneralToCodetree o #1 o recEnv) in SOME(simpPostSetContainer(optCont, Tuple{isVariant=false, fields=fields}, cDecs, filter)) end | _ => SOME(simpPostSetContainer(optCont, cGen, cDecs, filter)) end | simpGeneral (context as { enterAddr, nextAddress, reprocess, ...}) (BeginLoop{loop, arguments, ...}) = let val didReprocess = ! reprocess (* To see if we really need the loop first try simply binding the arguments and process it. It's often the case that if one or more arguments is a constant that the looping case will be eliminated. *) val withoutBeginLoop = simplify(mkEnv(List.map (Declar o #1) arguments, loop), context) fun foldLoop f n (Loop l) = f(l, n) | foldLoop f n (Newenv(_, exp)) = foldLoop f n exp | foldLoop f n (Cond(_, t, e)) = foldLoop f (foldLoop f n t) e | foldLoop f n (Handle {handler, ...}) = foldLoop f n handler | foldLoop f n (SetContainer{tuple, ...}) = foldLoop f n tuple | foldLoop _ n _ = n (* Check if the Loop instruction is there. This assumes that these are the only tail-recursive cases. *) val hasLoop = foldLoop (fn _ => true) false in if not (hasLoop withoutBeginLoop) then SOME withoutBeginLoop else let (* Reset "reprocess". It may have been set in the withoutBeginLoop that's not the code we're going to return. *) val () = reprocess := didReprocess (* We need the BeginLoop. Create new addresses for the arguments. *) fun declArg({addr, value, use, ...}, typ) = let val newAddr = nextAddress() in enterAddr(addr, (EnvGenLoad(LoadLocal newAddr), EnvSpecNone)); ({addr = newAddr, value = simplify(value, context), use = use }, typ) end (* Now look to see if the (remaining) loops have any arguments that do not change. Do this after processing because we could be eliminating other loops that may change the arguments. *) val declArgs = map declArg arguments val beginBody = simplify(loop, context) local fun argsMatch((Extract (LoadLocal argNo), _), ({addr, ...}, _)) = argNo = addr | argsMatch _ = false fun checkLoopArgs(loopArgs, checks) = let fun map3(loopA :: loopArgs, decA :: decArgs, checkA :: checkArgs) = (argsMatch(loopA, decA) andalso checkA) :: map3(loopArgs, decArgs, checkArgs) | map3 _ = [] in map3(loopArgs, declArgs, checks) end in val checkList = foldLoop checkLoopArgs (map (fn _ => true) arguments) beginBody end in if List.exists (fn l => l) checkList then let (* Turn the original arguments into bindings. *) local fun argLists(true, (arg, _), (tArgs, fArgs)) = (Declar arg :: tArgs, fArgs) | argLists(false, arg, (tArgs, fArgs)) = (tArgs, arg :: fArgs) in val (unchangedArgs, filteredDeclArgs) = ListPair.foldrEq argLists ([], []) (checkList, declArgs) end fun changeLoops (Loop loopArgs) = let val newArgs = ListPair.foldrEq(fn (false, arg, l) => arg :: l | (true, _, l) => l) [] (checkList, loopArgs) in Loop newArgs end | changeLoops(Newenv(decs, exp)) = Newenv(decs, changeLoops exp) | changeLoops(Cond(i, t, e)) = Cond(i, changeLoops t, changeLoops e) | changeLoops(Handle{handler, exp, exPacketAddr}) = Handle{handler=changeLoops handler, exp=exp, exPacketAddr=exPacketAddr} | changeLoops(SetContainer{tuple, container, filter}) = SetContainer{tuple=changeLoops tuple, container=container, filter=filter} | changeLoops code = code val beginBody = simplify(changeLoops loop, context) (* Reprocess because we've lost any special part from the arguments that haven't changed. *) val () = reprocess := true in SOME(mkEnv(unchangedArgs, BeginLoop {loop=beginBody, arguments=filteredDeclArgs})) end else SOME(BeginLoop {loop=beginBody, arguments=declArgs}) end end | simpGeneral context (TagTest{test, tag, maxTag}) = ( case simplify(test, context) of Constnt(testResult, _) => if isShort testResult andalso toShort testResult = tag then SOME CodeTrue else SOME CodeFalse | sTest => SOME(TagTest{test=sTest, tag=tag, maxTag=maxTag}) ) | simpGeneral context (LoadOperation{kind, address}) = let (* Try to move constants out of the index. *) val (genAddress, RevList decAddress) = simpAddress(address, getMultiplier kind, context) (* If the base address and index are constant and this is an immutable load we can do this at compile time. *) val result = case (genAddress, kind) of ({base=Constnt(baseAddr, _), index=NONE, offset}, LoadStoreMLWord _) => if isShort baseAddr then LoadOperation{kind=kind, address=genAddress} else let (* Ignore the "isImmutable" flag and look at the immutable status of the memory. Check that this is a word object and that the offset is within range. The code for Vector.sub, for example, raises an exception if the index is out of range but still generates the (unreachable) indexing code. *) val addr = toAddress baseAddr val wordOffset = offset div RunCall.bytesPerWord in if isMutable addr orelse not(isWords addr) orelse wordOffset >= length addr then LoadOperation{kind=kind, address=genAddress} else Constnt(toMachineWord(loadWord(addr, wordOffset)), []) end | ({base=Constnt(baseAddr, _), index=NONE, offset}, LoadStoreMLByte _) => if isShort baseAddr then LoadOperation{kind=kind, address=genAddress} else let val addr = toAddress baseAddr val wordOffset = offset div RunCall.bytesPerWord in if isMutable addr orelse not(isBytes addr) orelse wordOffset >= length addr then LoadOperation{kind=kind, address=genAddress} else Constnt(toMachineWord(loadByte(addr, offset)), []) end | ({base=Constnt(baseAddr, _), index=NONE, offset}, LoadStoreUntaggedUnsigned) => if isShort baseAddr then LoadOperation{kind=kind, address=genAddress} else let val addr = toAddress baseAddr (* We don't currently have loadWordUntagged in Address but it's only ever used to load the string length word so we can use that. *) in if isMutable addr orelse not(isBytes addr) orelse offset <> 0w0 then LoadOperation{kind=kind, address=genAddress} else Constnt(toMachineWord(String.size(RunCall.unsafeCast addr)), []) end | _ => LoadOperation{kind=kind, address=genAddress} in SOME(mkEnv(List.rev decAddress, result)) end | simpGeneral context (StoreOperation{kind, address, value}) = let val (genAddress, decAddress) = simpAddress(address, getMultiplier kind, context) val (genValue, RevList decValue, _) = simpSpecial(value, context, decAddress) in SOME(mkEnv(List.rev decValue, StoreOperation{kind=kind, address=genAddress, value=genValue})) end | simpGeneral (context as {reprocess, ...}) (BlockOperation{kind, sourceLeft, destRight, length}) = let val multiplier = case kind of BlockOpMove{isByteMove=false} => RunCall.bytesPerWord | BlockOpMove{isByteMove=true} => 0w1 | BlockOpEqualByte => 0w1 | BlockOpCompareByte => 0w1 val (genSrcAddress, RevList decSrcAddress) = simpAddress(sourceLeft, multiplier, context) val (genDstAddress, RevList decDstAddress) = simpAddress(destRight, multiplier, context) val (genLength, RevList decLength, _) = simpSpecial(length, context, RevList []) (* If we have a short length move we're better doing it as a sequence of loads and stores. This is particularly useful with string concatenation. Small here means three or less. Four and eight byte moves are handled as single instructions in the code-generator provided the alignment is correct. *) val shortLength = case genLength of Constnt(lenConst, _) => if isShort lenConst then let val l = toShort lenConst in if l <= 0w3 then SOME l else NONE end else NONE | _ => NONE val combinedDecs = List.rev decSrcAddress @ List.rev decDstAddress @ List.rev decLength val operation = case (shortLength, kind) of (SOME length, BlockOpMove{isByteMove}) => let val _ = reprocess := true (* Frequently the source will be a constant. *) val {base=baseSrc, index=indexSrc, offset=offsetSrc} = genSrcAddress and {base=baseDst, index=indexDst, offset=offsetDst} = genDstAddress (* We don't know if the source is immutable but the destination definitely isn't *) val moveKind = if isByteMove then LoadStoreMLByte{isImmutable=false} else LoadStoreMLWord{isImmutable=false} fun makeMoves offset = if offset = length then [] else NullBinding( StoreOperation{kind=moveKind, address={base=baseDst, index=indexDst, offset=offsetDst+offset*multiplier}, value=LoadOperation{kind=moveKind, address={base=baseSrc, index=indexSrc, offset=offsetSrc+offset*multiplier}}}) :: makeMoves(offset+0w1) in mkEnv(combinedDecs @ makeMoves 0w0, CodeZero (* unit result *)) end | (SOME length, BlockOpEqualByte) => (* Comparing with the null string and up to 3 characters. *) let val {base=baseSrc, index=indexSrc, offset=offsetSrc} = genSrcAddress and {base=baseDst, index=indexDst, offset=offsetDst} = genDstAddress val moveKind = LoadStoreMLByte{isImmutable=false} (* Build andalso tree to check each byte. For the null string this simply returns "true". *) fun makeComparison offset = if offset = length then CodeTrue else Cond( Binary{oper=WordComparison{test=TestEqual, isSigned=false}, arg1=LoadOperation{kind=moveKind, address={base=baseSrc, index=indexSrc, offset=offsetSrc+offset*multiplier}}, arg2=LoadOperation{kind=moveKind, address={base=baseDst, index=indexDst, offset=offsetDst+offset*multiplier}}}, makeComparison(offset+0w1), CodeFalse) in mkEnv(combinedDecs, makeComparison 0w0) end | _ => mkEnv(combinedDecs, BlockOperation{kind=kind, sourceLeft=genSrcAddress, destRight=genDstAddress, length=genLength}) in SOME operation end | simpGeneral (context as {enterAddr, nextAddress, ...}) (Handle{exp, handler, exPacketAddr}) = let (* We need to make a new binding for the exception packet. *) val expBody = simplify(exp, context) val newAddr = nextAddress() val () = enterAddr(exPacketAddr, (EnvGenLoad(LoadLocal newAddr), EnvSpecNone)) val handleBody = simplify(handler, context) in SOME(Handle{exp=expBody, handler=handleBody, exPacketAddr=newAddr}) end | simpGeneral _ _ = NONE (* Where we have an Indirect or Eval we want the argument as either a tuple or an inline function respectively if that's possible. Getting that also involves various other cases as well. Because a binding may later be used in such a context we treat any binding in that way as well. *) and simpSpecial (Extract ext, { lookupAddr, ...}, tailDecs) = let val (gen, spec) = lookupAddr ext in (envGeneralToCodetree gen, tailDecs, spec) end | simpSpecial (Newenv envArgs, context, tailDecs) = simpNewenv(envArgs, context, tailDecs) | simpSpecial (Lambda lambda, context, tailDecs) = let val (gen, spec) = simpLambda(lambda, context, NONE, NONE) in (Lambda gen, tailDecs, spec) end | simpSpecial (Eval {function, argList, resultType}, context, tailDecs) = simpFunctionCall(function, argList, resultType, context, tailDecs) | simpSpecial (Unary{oper, arg1}, context, tailDecs) = simpUnary(oper, arg1, context, tailDecs) | simpSpecial (Binary{oper, arg1, arg2}, context, tailDecs) = simpBinary(oper, arg1, arg2, context, tailDecs) | simpSpecial (Arbitrary{oper=ArbCompare test, shortCond, arg1, arg2, longCall}, context, tailDecs) = simpArbitraryCompare(test, shortCond, arg1, arg2, longCall, context, tailDecs) | simpSpecial (Arbitrary{oper=ArbArith arith, shortCond, arg1, arg2, longCall}, context, tailDecs) = simpArbitraryArith(arith, shortCond, arg1, arg2, longCall, context, tailDecs) | simpSpecial (AllocateWordMemory{numWords, flags, initial}, context, tailDecs) = simpAllocateWordMemory(numWords, flags, initial, context, tailDecs) | simpSpecial (Cond(condTest, condThen, condElse), context, tailDecs) = simpIfThenElse(condTest, condThen, condElse, context, tailDecs) | simpSpecial (Tuple { fields, isVariant }, context, tailDecs) = simpTuple(fields, isVariant, context, tailDecs) | simpSpecial (Indirect{ base, offset, indKind }, context, tailDecs) = simpFieldSelect(base, offset, indKind, context, tailDecs) | simpSpecial (c: codetree, s: simpContext, tailDecs): codetree * revlist * envSpecial = let (* Anything else - copy it and then split it into the fields. *) fun split(Newenv(l, e), RevList tailDecs) = (* Pull off bindings. *) split (e, RevList(List.rev l @ tailDecs)) | split(Constnt(m, p), tailDecs) = (Constnt(m, p), tailDecs, findInline p) | split(c, tailDecs) = (c, tailDecs, EnvSpecNone) in split(simplify(c, s), tailDecs) end (* Process a Newenv. We need to add the bindings to the context. *) and simpNewenv((envDecs: codeBinding list, envExp), context as { enterAddr, nextAddress, reprocess, ...}, tailDecs): codetree * revlist * envSpecial = let fun copyDecs ([], decs) = simpSpecial(envExp, context, decs) (* End of the list - process the result expression. *) | copyDecs ((Declar{addr, value, ...} :: vs), decs) = ( case simpSpecial(value, context, decs) of (* If this raises an exception stop here. *) vBinding as (Raise _, _, _) => vBinding | vBinding => let (* Add the declaration to the table. *) val (optV, dec) = makeNewDecl(vBinding, context) val () = enterAddr(addr, optV) in copyDecs(vs, dec) end ) | copyDecs(NullBinding v :: vs, decs) = (* Not a binding - process this and the rest.*) ( case simpSpecial(v, context, decs) of (* If this raises an exception stop here. *) vBinding as (Raise _, _, _) => vBinding | (cGen, RevList cDecs, _) => copyDecs(vs, RevList(NullBinding cGen :: cDecs)) ) | copyDecs(RecDecs mutuals :: vs, RevList decs) = (* Mutually recursive declarations. Any of the declarations may refer to any of the others. They should all be lambdas. The front end generates functions with more than one argument (either curried or tupled) as pairs of mutually recursive functions. The main function body takes its arguments on the stack (or in registers) and the auxiliary inline function, possibly nested, takes the tupled or curried arguments and calls it. If the main function is recursive it will first call the inline function which is why the pair are mutually recursive. As far as possible we want to use the main function since that uses the least memory. Specifically, if the function recurses we want the recursive call to pass all the arguments if it can. *) let (* Reorder the function so the explicitly-inlined ones come first. Their code can then be inserted into the main functions. *) local val (inlines, nonInlines) = List.partition ( fn {lambda = { isInline=Inline, ...}, ... } => true | _ => false) mutuals in val orderedDecs = inlines @ nonInlines end (* Go down the functions creating new addresses for them and entering them in the table. *) val addresses = map (fn {addr, ... } => let val decAddr = nextAddress() in enterAddr (addr, (EnvGenLoad(LoadLocal decAddr), EnvSpecNone)); decAddr end) orderedDecs fun processFunction({ lambda, addr, ... }, newAddr) = let val (gen, spec) = simpLambda(lambda, context, SOME addr, SOME newAddr) (* Update the entry in the table to include any inlineable function. *) val () = enterAddr (addr, (EnvGenLoad (LoadLocal newAddr), spec)) in {addr=newAddr, lambda=gen, use=[]} end val rlist = ListPair.map processFunction (orderedDecs, addresses) in (* and put these declarations onto the list. *) copyDecs(vs, RevList(List.rev(partitionMutualBindings(RecDecs rlist)) @ decs)) end | copyDecs (Container{addr, size, setter, ...} :: vs, RevList decs) = let (* Enter the new address immediately - it's needed in the setter. *) val decAddr = nextAddress() val () = enterAddr (addr, (EnvGenLoad(LoadLocal decAddr), EnvSpecNone)) val (setGen, RevList setDecs, _) = simpSpecial(setter, context, RevList []) in (* If we have inline expanded a function that sets the container we're better off eliminating the container completely. *) case setGen of SetContainer { tuple, filter, container } => let (* Check the container we're setting is the address we've made for it. *) val _ = (case container of Extract(LoadLocal a) => a = decAddr | _ => false) orelse raise InternalError "copyDecs: Container/SetContainer" val newDecAddr = nextAddress() val () = enterAddr (addr, (EnvGenLoad(LoadLocal newDecAddr), EnvSpecNone)) val tupleAddr = nextAddress() val tupleDec = Declar{addr=tupleAddr, use=[], value=tuple} val tupleLoad = mkLoadLocal tupleAddr val resultTuple = BoolVector.foldri(fn (i, true, l) => mkInd(i, tupleLoad) :: l | (_, false, l) => l) [] filter val _ = List.length resultTuple = size orelse raise InternalError "copyDecs: Container/SetContainer size" val containerDec = Declar{addr=newDecAddr, use=[], value=mkTuple resultTuple} (* TODO: We're replacing a container with what is notionally a tuple on the heap. It should be optimised away as a result of a further pass but we currently have indirections from a container for these. On the native platforms that doesn't matter but on 32-in-64 indirecting from the heap and from the stack are different. *) val _ = reprocess := true in copyDecs(vs, RevList(containerDec :: tupleDec :: setDecs @ decs)) end | _ => let (* The setDecs could refer the container itself if we've optimised this with simpPostSetContainer so we must include them within the setter and not lift them out. *) val dec = Container{addr=decAddr, use=[], size=size, setter=mkEnv(List.rev setDecs, setGen)} in copyDecs(vs, RevList(dec :: decs)) end end in copyDecs(envDecs, tailDecs) end (* Prepares a binding for entry into a look-up table. Returns the entry to put into the table together with any bindings that must be made. If the general part of the optVal is a constant we can just put the constant in the table. If it is a load (Extract) it is just renaming an existing entry so we can return it. Otherwise we have to make a new binding and return a load (Extract) entry for it. *) and makeNewDecl((Constnt w, RevList decs, spec), _) = ((EnvGenConst w, spec), RevList decs) (* No need to create a binding for a constant. *) | makeNewDecl((Extract ext, RevList decs, spec), _) = ((EnvGenLoad ext, spec), RevList decs) (* Binding is simply giving a new name to a variable - can ignore this declaration. *) | makeNewDecl((gen, RevList decs, spec), { nextAddress, ...}) = let (* Create a binding for this value. *) val newAddr = nextAddress() in ((EnvGenLoad(LoadLocal newAddr), spec), RevList(mkDec(newAddr, gen) :: decs)) end and simpLambda({body, isInline, name, argTypes, resultType, closure, localCount, ...}, { lookupAddr, reprocess, ... }, myOldAddrOpt, myNewAddrOpt) = let (* A new table for the new function. *) val oldAddrTab = Array.array (localCount, NONE) val optClosureList = makeClosure() val isNowRecursive = ref false local fun localOldAddr (LoadLocal addr) = valOf(Array.sub(oldAddrTab, addr)) | localOldAddr (ext as LoadArgument _) = (EnvGenLoad ext, EnvSpecNone) | localOldAddr (ext as LoadRecursive) = (EnvGenLoad ext, EnvSpecNone) | localOldAddr (LoadClosure addr) = let val oldEntry = List.nth(closure, addr) (* If the entry in the closure is our own address this is recursive. *) fun isRecursive(EnvGenLoad(LoadLocal a), SOME b) = if a = b then (isNowRecursive := true; true) else false | isRecursive _ = false in if isRecursive(EnvGenLoad oldEntry, myOldAddrOpt) then (EnvGenLoad LoadRecursive, EnvSpecNone) else let val newEntry = lookupAddr oldEntry val makeClosure = addToClosure optClosureList fun convertResult(genEntry, specEntry) = (* If after looking up the entry we get our new address it's recursive. *) if isRecursive(genEntry, myNewAddrOpt) then (EnvGenLoad LoadRecursive, EnvSpecNone) else let val newGeneral = case genEntry of EnvGenLoad ext => EnvGenLoad(makeClosure ext) | EnvGenConst w => EnvGenConst w (* Have to modify the environment here so that if we look up free variables we add them to the closure. *) fun convertEnv env args = convertResult(env args) val newSpecial = case specEntry of EnvSpecTuple(size, env) => EnvSpecTuple(size, convertEnv env) | EnvSpecInlineFunction(spec, env) => EnvSpecInlineFunction(spec, convertEnv env) | EnvSpecUnary _ => EnvSpecNone (* Don't pass this in *) | EnvSpecBinary _ => EnvSpecNone (* Don't pass this in *) | EnvSpecNone => EnvSpecNone in (newGeneral, newSpecial) end in convertResult newEntry end end and setTab (index, v) = Array.update (oldAddrTab, index, SOME v) in val newAddressAllocator = ref 0 fun mkAddr () = ! newAddressAllocator before newAddressAllocator := ! newAddressAllocator + 1 val newCode = simplify (body, { enterAddr = setTab, lookupAddr = localOldAddr, nextAddress=mkAddr, reprocess = reprocess }) end val closureAfterOpt = extractClosure optClosureList val localCount = ! newAddressAllocator (* If we have mutually recursive "small" functions we may turn them into recursive functions. We have to remove the "small" status from them to prevent them from being expanded inline anywhere else. The optimiser may turn them back into "small" functions if the recursion is actually tail-recursion. *) val isNowInline = case isInline of Inline => if ! isNowRecursive then NonInline else Inline | NonInline => NonInline (* Clean up the function body at this point if it could be inlined. There are examples where failing to do this can blow up. This can be the result of creating both a general and special function inside an inline function. *) val cleanBody = case isNowInline of NonInline => newCode | _ => REMOVE_REDUNDANT.cleanProc(newCode, [UseExport], LoadClosure, localCount) val copiedLambda: lambdaForm = { body = cleanBody, isInline = isNowInline, name = name, closure = closureAfterOpt, argTypes = argTypes, resultType = resultType, localCount = localCount, recUse = [] } val inlineCode = case isNowInline of NonInline => EnvSpecNone | _ => EnvSpecInlineFunction(copiedLambda, fn addr => (EnvGenLoad(List.nth(closureAfterOpt, addr)), EnvSpecNone)) in ( copiedLambda, inlineCode ) end and simpFunctionCall(function, argList, resultType, context as { reprocess, ...}, tailDecs) = let (* Function call - This may involve inlining the function. *) (* Get the function to be called and see if it is inline or a lambda expression. *) val (genFunct, decsFunct, specFunct) = simpSpecial(function, context, tailDecs) (* We have to make a special check here that we are not passing in the function we are trying to expand. This could result in an infinitely recursive expansion. It is only going to happen in very special circumstances such as a definition of the Y combinator. If we see that we don't attempt to expand inline. It could be embedded in a tuple or the closure of a function as well as passed directly. *) val isRecursiveArg = case function of Extract extOrig => let fun containsFunction(Extract thisArg, v) = (v orelse thisArg = extOrig, FOLD_DESCEND) | containsFunction(Lambda{closure, ...}, v) = (* Only the closure, not the body *) (foldl (fn (c, w) => foldtree containsFunction w (Extract c)) v closure, FOLD_DONT_DESCEND) | containsFunction(Eval _, v) = (v, FOLD_DONT_DESCEND) (* OK if it's called *) | containsFunction(_, v) = (v, FOLD_DESCEND) in List.exists(fn (c, _) => foldtree containsFunction false c) argList end | _ => false in case (specFunct, genFunct, isRecursiveArg) of (EnvSpecInlineFunction({body=lambdaBody, localCount, argTypes, ...}, functEnv), _, false) => let val _ = List.length argTypes = List.length argList orelse raise InternalError "simpFunctionCall: argument mismatch" val () = reprocess := true (* If we expand inline we have to reprocess *) and { nextAddress, reprocess, ...} = context (* Expand a function inline, either one marked explicitly to be inlined or one detected as "small". *) (* Calling inline proc or a lambda expression which is just called. The function is replaced with a block containing declarations of the parameters. We need a new table here because the addresses we use to index it are the addresses which are local to the function. New addresses are created in the range of the surrounding function. *) val localVec = Array.array(localCount, NONE) local fun processArgs([], bindings) = ([], bindings) | processArgs((arg, _)::args, bindings) = let val (thisArg, newBindings) = makeNewDecl(simpSpecial(arg, context, bindings), context) val (otherArgs, resBindings) = processArgs(args, newBindings) in (thisArg::otherArgs, resBindings) end val (params, bindings) = processArgs(argList, decsFunct) val paramVec = Vector.fromList params in fun getParameter n = Vector.sub(paramVec, n) (* Bindings necessary for the arguments *) val copiedArgs = bindings end local fun localOldAddr(LoadLocal addr) = valOf(Array.sub(localVec, addr)) | localOldAddr(LoadArgument addr) = getParameter addr | localOldAddr(LoadClosure closureEntry) = functEnv closureEntry | localOldAddr LoadRecursive = raise InternalError "localOldAddr: LoadRecursive" fun setTabForInline (index, v) = Array.update (localVec, index, SOME v) val lambdaContext = { lookupAddr=localOldAddr, enterAddr=setTabForInline, nextAddress=nextAddress, reprocess = reprocess } in val (cGen, cDecs, cSpec) = simpSpecial(lambdaBody,lambdaContext, copiedArgs) end in (cGen, cDecs, cSpec) end | (_, gen as Constnt _, _) => (* Not inlinable - constant function. *) let val copiedArgs = map (fn (arg, argType) => (simplify(arg, context), argType)) argList val evCopiedCode = Eval {function = gen, argList = copiedArgs, resultType=resultType} in (evCopiedCode, decsFunct, EnvSpecNone) end | (_, gen, _) => (* Anything else. *) let val copiedArgs = map (fn (arg, argType) => (simplify(arg, context), argType)) argList val evCopiedCode = Eval {function = gen, argList = copiedArgs, resultType=resultType} in (evCopiedCode, decsFunct, EnvSpecNone) end end (* Special processing for the current builtIn1 operations. *) (* Constant folding for built-ins. These ought to be type-correct i.e. we should have tagged values in some cases and addresses in others. However there may be run-time tests that would ensure type-correctness and we can't be sure that they will always be folded at compile-time. e.g. we may have if isShort c then shortOp c else longOp c If c is a constant then we may try to fold both the shortOp and the longOp and one of these will be type-incorrect although never executed at run-time. *) and simpUnary(oper, arg1, context as { reprocess, ...}, tailDecs) = let val (genArg1, decArg1, specArg1) = simpSpecial(arg1, context, tailDecs) in case (oper, genArg1) of (NotBoolean, Constnt(v, _)) => ( reprocess := true; (if isShort v andalso toShort v = 0w0 then CodeTrue else CodeFalse, decArg1, EnvSpecNone) ) | (NotBoolean, genArg1) => ( (* NotBoolean: This can be the result of using Bool.not but more usually occurs as a result of other code. We don't have TestNotEqual or IsAddress so both of these use NotBoolean with TestEqual and IsTagged. Also we can insert a NotBoolean as a result of a Cond. We try to eliminate not(not a) and to push other NotBooleans down to a point where a boolean is tested. *) case specArg1 of EnvSpecUnary(NotBoolean, originalArg) => ( (* not(not a) - Eliminate. *) reprocess := true; (originalArg, decArg1, EnvSpecNone) ) | _ => (* Otherwise pass this on. It is also extracted in a Cond. *) (Unary{oper=NotBoolean, arg1=genArg1}, decArg1, EnvSpecUnary(NotBoolean, genArg1)) ) | (IsTaggedValue, Constnt(v, _)) => ( reprocess := true; (if isShort v then CodeTrue else CodeFalse, decArg1, EnvSpecNone) ) | (IsTaggedValue, genArg1) => ( (* We use this to test for nil values and if we have constructed a record (or possibly a function) it can't be null. *) case specArg1 of EnvSpecTuple _ => (CodeFalse, decArg1, EnvSpecNone) before reprocess := true | EnvSpecInlineFunction _ => (CodeFalse, decArg1, EnvSpecNone) before reprocess := true | _ => (Unary{oper=oper, arg1=genArg1}, decArg1, EnvSpecNone) ) | (MemoryCellLength, Constnt(v, _)) => ( reprocess := true; (if isShort v then CodeZero else Constnt(toMachineWord(Address.length(toAddress v)), []), decArg1, EnvSpecNone) ) | (MemoryCellFlags, Constnt(v, _)) => ( reprocess := true; (if isShort v then CodeZero else Constnt(toMachineWord(Address.flags(toAddress v)), []), decArg1, EnvSpecNone) ) | (LongWordToTagged, Constnt(v, _)) => ( reprocess := true; (Constnt(toMachineWord(Word.fromLargeWord(RunCall.unsafeCast v)), []), decArg1, EnvSpecNone) ) | (LongWordToTagged, genArg1) => ( (* If we apply LongWordToTagged to an argument we have created with UnsignedToLongWord we can return the original argument. *) case specArg1 of EnvSpecUnary(UnsignedToLongWord, originalArg) => ( reprocess := true; (originalArg, decArg1, EnvSpecNone) ) | _ => (Unary{oper=LongWordToTagged, arg1=genArg1}, decArg1, EnvSpecNone) ) | (SignedToLongWord, Constnt(v, _)) => ( reprocess := true; (Constnt(toMachineWord(Word.toLargeWordX(RunCall.unsafeCast v)), []), decArg1, EnvSpecNone) ) | (UnsignedToLongWord, Constnt(v, _)) => ( reprocess := true; (Constnt(toMachineWord(Word.toLargeWord(RunCall.unsafeCast v)), []), decArg1, EnvSpecNone) ) | (UnsignedToLongWord, genArg1) => (* Add the operation as the special entry. It can then be recognised by LongWordToTagged. *) (Unary{oper=oper, arg1=genArg1}, decArg1, EnvSpecUnary(UnsignedToLongWord, genArg1)) | _ => (Unary{oper=oper, arg1=genArg1}, decArg1, EnvSpecNone) end and simpBinary(oper, arg1, arg2, context as {reprocess, ...}, tailDecs) = let val (genArg1, decArg1, _ (*specArg1*)) = simpSpecial(arg1, context, tailDecs) val (genArg2, decArgs, _ (*specArg2*)) = simpSpecial(arg2, context, decArg1) in case (oper, genArg1, genArg2) of (WordComparison{test, isSigned}, Constnt(v1, _), Constnt(v2, _)) => - if (case test of TestEqual => false | _ => not(isShort v1) orelse not(isShort v2)) + if not(isShort v1) orelse not(isShort v2) (* E.g. arbitrary precision on unreachable path. *) then (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) else let val () = reprocess := true val testResult = case (test, isSigned) of (* TestEqual can be applied to addresses. *) - (TestEqual, _) => RunCall.pointerEq(v1, v2) + (TestEqual, _) => toShort v1 = toShort v2 | (TestLess, false) => toShort v1 < toShort v2 | (TestLessEqual, false) => toShort v1 <= toShort v2 | (TestGreater, false) => toShort v1 > toShort v2 | (TestGreaterEqual, false) => toShort v1 >= toShort v2 | (TestLess, true) => toFix v1 < toFix v2 | (TestLessEqual, true) => toFix v1 <= toFix v2 | (TestGreater, true) => toFix v1 > toFix v2 | (TestGreaterEqual, true) => toFix v1 >= toFix v2 | (TestUnordered, _) => raise InternalError "WordComparison: TestUnordered" in (if testResult then CodeTrue else CodeFalse, decArgs, EnvSpecNone) end + + | (PointerEq, Constnt(v1, _), Constnt(v2, _)) => + ( + reprocess := true; + (if RunCall.pointerEq(v1, v2) then CodeTrue else CodeFalse, decArgs, EnvSpecNone) + ) | (FixedPrecisionArith arithOp, Constnt(v1, _), Constnt(v2, _)) => if not(isShort v1) orelse not(isShort v2) then (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) else let val () = reprocess := true val v1S = toFix v1 and v2S = toFix v2 fun asConstnt v = Constnt(toMachineWord v, []) val raiseOverflow = Raise(Constnt(toMachineWord Overflow, [])) val raiseDiv = Raise(Constnt(toMachineWord Div, [])) (* ?? There's usually an explicit test. *) val resultCode = case arithOp of ArithAdd => (asConstnt(v1S+v2S) handle Overflow => raiseOverflow) | ArithSub => (asConstnt(v1S-v2S) handle Overflow => raiseOverflow) | ArithMult => (asConstnt(v1S*v2S) handle Overflow => raiseOverflow) | ArithQuot => (asConstnt(FixedInt.quot(v1S,v2S)) handle Overflow => raiseOverflow | Div => raiseDiv) | ArithRem => (asConstnt(FixedInt.rem(v1S,v2S)) handle Overflow => raiseOverflow | Div => raiseDiv) | ArithDiv => (asConstnt(FixedInt.div(v1S,v2S)) handle Overflow => raiseOverflow | Div => raiseDiv) | ArithMod => (asConstnt(FixedInt.mod(v1S,v2S)) handle Overflow => raiseOverflow | Div => raiseDiv) in (resultCode, decArgs, EnvSpecNone) end (* Addition and subtraction of zero. These can arise as a result of inline expansion of more general functions. *) | (FixedPrecisionArith ArithAdd, arg1, Constnt(v2, _)) => if isShort v2 andalso toShort v2 = 0w0 then (arg1, decArgs, EnvSpecNone) else (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) | (FixedPrecisionArith ArithAdd, Constnt(v1, _), arg2) => if isShort v1 andalso toShort v1 = 0w0 then (arg2, decArgs, EnvSpecNone) else (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) | (FixedPrecisionArith ArithSub, arg1, Constnt(v2, _)) => if isShort v2 andalso toShort v2 = 0w0 then (arg1, decArgs, EnvSpecNone) else (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) | (WordArith arithOp, Constnt(v1, _), Constnt(v2, _)) => if not(isShort v1) orelse not(isShort v2) then (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) else let val () = reprocess := true val v1S = toShort v1 and v2S = toShort v2 fun asConstnt v = Constnt(toMachineWord v, []) val resultCode = case arithOp of ArithAdd => asConstnt(v1S+v2S) | ArithSub => asConstnt(v1S-v2S) | ArithMult => asConstnt(v1S*v2S) | ArithQuot => raise InternalError "WordArith: ArithQuot" | ArithRem => raise InternalError "WordArith: ArithRem" | ArithDiv => asConstnt(v1S div v2S) | ArithMod => asConstnt(v1S mod v2S) in (resultCode, decArgs, EnvSpecNone) end | (WordArith ArithAdd, arg1, Constnt(v2, _)) => if isShort v2 andalso toShort v2 = 0w0 then (arg1, decArgs, EnvSpecNone) else (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) | (WordArith ArithAdd, Constnt(v1, _), arg2) => if isShort v1 andalso toShort v1 = 0w0 then (arg2, decArgs, EnvSpecNone) else (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) | (WordArith ArithSub, arg1, Constnt(v2, _)) => if isShort v2 andalso toShort v2 = 0w0 then (arg1, decArgs, EnvSpecNone) else (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) | (WordLogical logOp, Constnt(v1, _), Constnt(v2, _)) => if not(isShort v1) orelse not(isShort v2) then (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) else let val () = reprocess := true val v1S = toShort v1 and v2S = toShort v2 fun asConstnt v = Constnt(toMachineWord v, []) val resultCode = case logOp of LogicalAnd => asConstnt(Word.andb(v1S,v2S)) | LogicalOr => asConstnt(Word.orb(v1S,v2S)) | LogicalXor => asConstnt(Word.xorb(v1S,v2S)) in (resultCode, decArgs, EnvSpecNone) end | (WordLogical logop, arg1, Constnt(v2, _)) => (* Return the zero if we are anding with zero otherwise the original arg *) if isShort v2 andalso toShort v2 = 0w0 then (case logop of LogicalAnd => CodeZero | _ => arg1, decArgs, EnvSpecNone) else (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) | (WordLogical logop, Constnt(v1, _), arg2) => if isShort v1 andalso toShort v1 = 0w0 then (case logop of LogicalAnd => CodeZero | _ => arg2, decArgs, EnvSpecNone) else (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) (* TODO: Constant folding of shifts. *) | _ => (Binary{oper=oper, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) end (* Arbitrary precision operations. This is a sort of mixture of a built-in and a conditional. *) and simpArbitraryCompare(TestEqual, shortCond, arg1, arg2, longCall, context, tailDecs) = (* Equality is a special case and is only there to ensure that it is not accidentally converted into an indexed case further down. We must leave it as it is. *) let val (genCond, decCond, _ (*specArg1*)) = simpSpecial(shortCond, context, tailDecs) val (genArg1, decArg1, _ (*specArg1*)) = simpSpecial(arg1, context, decCond) val (genArg2, decArgs, _ (*specArg2*)) = simpSpecial(arg2, context, decArg1) in case (genArg1, genArg2) of (Constnt(v1, _), Constnt(v2, _)) => let val a1: LargeInt.int = RunCall.unsafeCast v1 and a2: LargeInt.int = RunCall.unsafeCast v2 in (if a1 = a2 then CodeTrue else CodeFalse, decArgs, EnvSpecNone) end | _ => (Arbitrary{oper=ArbCompare TestEqual, shortCond=genCond, arg1=genArg1, arg2=genArg2, longCall=simplify(longCall, context)}, decArgs, EnvSpecNone) end | simpArbitraryCompare(test, shortCond, arg1, arg2, longCall, context as {reprocess, ...}, tailDecs) = let val (genCond, decCond, _ (*specArg1*)) = simpSpecial(shortCond, context, tailDecs) val (genArg1, decArg1, _ (*specArg1*)) = simpSpecial(arg1, context, decCond) val (genArg2, decArgs, _ (*specArg2*)) = simpSpecial(arg2, context, decArg1) val posFlags = Address.F_bytes and negFlags = Word8.orb(Address.F_bytes, Address.F_negative) in (* Fold any constant/constant operations but more importantly, if we have variable/constant operations where the constant is short we can avoid using the full arbitrary precision call by just looking at the sign bit. *) case (genCond, genArg1, genArg2) of (_, Constnt(v1, _), Constnt(v2, _)) => let val a1: LargeInt.int = RunCall.unsafeCast v1 and a2: LargeInt.int = RunCall.unsafeCast v2 val testResult = case test of TestLess => a1 < a2 | TestGreater => a1 > a2 | TestLessEqual => a1 <= a2 | TestGreaterEqual => a1 >= a2 | _ => raise InternalError "simpArbitraryCompare: Unimplemented function" in (if testResult then CodeTrue else CodeFalse, decArgs, EnvSpecNone) end | (Constnt(c1, _), _, _) => if isShort c1 andalso toShort c1 = 0w0 then (* One argument is definitely long - generate the long form. *) (Binary{oper=WordComparison{test=test, isSigned=true}, arg1=simplify(longCall, context), arg2=CodeZero}, decArgs, EnvSpecNone) else (* Both arguments are short. That should mean they're constants. *) (Binary{oper=WordComparison{test=test, isSigned=true}, arg1=genArg1, arg2=genArg2}, decArgs, EnvSpecNone) before reprocess := true | (_, genArg1, cArg2 as Constnt _) => let (* The constant must be short otherwise the test would be false. *) val isNeg = case test of TestLess => true | TestLessEqual => true | _ => false (* Translate i < c into if isShort i then toShort i < c else isNegative i *) val newCode = Cond(Unary{oper=BuiltIns.IsTaggedValue, arg1=genArg1}, Binary { oper = BuiltIns.WordComparison{test=test, isSigned=true}, arg1 = genArg1, arg2 = cArg2 }, Binary { oper = BuiltIns.WordComparison{test=TestEqual, isSigned=false}, arg1=Unary { oper = MemoryCellFlags, arg1=genArg1 }, arg2=Constnt(toMachineWord(if isNeg then negFlags else posFlags), [])} ) in (newCode, decArgs, EnvSpecNone) end | (_, cArg1 as Constnt _, genArg2) => let (* We're testing c < i so the test is if isShort i then c < toShort i else isPositive i *) val isPos = case test of TestLess => true | TestLessEqual => true | _ => false val newCode = Cond(Unary{oper=BuiltIns.IsTaggedValue, arg1=genArg2}, Binary { oper = BuiltIns.WordComparison{test=test, isSigned=true}, arg1 = cArg1, arg2 = genArg2 }, Binary { oper = BuiltIns.WordComparison{test=TestEqual, isSigned=false}, arg1=Unary { oper = MemoryCellFlags, arg1=genArg2 }, arg2=Constnt(toMachineWord(if isPos then posFlags else negFlags), [])} ) in (newCode, decArgs, EnvSpecNone) end | _ => (Arbitrary{oper=ArbCompare test, shortCond=genCond, arg1=genArg1, arg2=genArg2, longCall=simplify(longCall, context)}, decArgs, EnvSpecNone) end and simpArbitraryArith(arith, shortCond, arg1, arg2, longCall, context, tailDecs) = let (* arg1 and arg2 are the arguments. shortCond is the condition that must be satisfied in order to use the short precision operation i.e. each argument must be short. *) val (genCond, decCond, _ (*specArg1*)) = simpSpecial(shortCond, context, tailDecs) val (genArg1, decArg1, _ (*specArg1*)) = simpSpecial(arg1, context, decCond) val (genArg2, decArgs, _ (*specArg2*)) = simpSpecial(arg2, context, decArg1) in case (genArg1, genArg2, genCond) of (Constnt(v1, _), Constnt(v2, _), _) => let val a1: LargeInt.int = RunCall.unsafeCast v1 and a2: LargeInt.int = RunCall.unsafeCast v2 (*val _ = print ("Fold arbitrary precision: " ^ PolyML.makestring(arith, a1, a2) ^ "\n")*) in case arith of ArithAdd => (Constnt(toMachineWord(a1+a2), []), decArgs, EnvSpecNone) | ArithSub => (Constnt(toMachineWord(a1-a2), []), decArgs, EnvSpecNone) | ArithMult => (Constnt(toMachineWord(a1*a2), []), decArgs, EnvSpecNone) | _ => raise InternalError "simpArbitraryArith: Unimplemented function" end | (_, _, Constnt(c1, _)) => if isShort c1 andalso toShort c1 = 0w0 then (* One argument is definitely long - generate the long form. *) (simplify(longCall, context), decArgs, EnvSpecNone) else (Arbitrary{oper=ArbArith arith, shortCond=genCond, arg1=genArg1, arg2=genArg2, longCall=simplify(longCall, context)}, decArgs, EnvSpecNone) | _ => (Arbitrary{oper=ArbArith arith, shortCond=genCond, arg1=genArg1, arg2=genArg2, longCall=simplify(longCall, context)}, decArgs, EnvSpecNone) end and simpAllocateWordMemory(numWords, flags, initial, context, tailDecs) = let val (genArg1, decArg1, _ (*specArg1*)) = simpSpecial(numWords, context, tailDecs) val (genArg2, decArg2, _ (*specArg2*)) = simpSpecial(flags, context, decArg1) val (genArg3, decArg3, _ (*specArg3*)) = simpSpecial(initial, context, decArg2) in (AllocateWordMemory{numWords=genArg1, flags=genArg2, initial=genArg3}, decArg3, EnvSpecNone) end (* Loads, stores and block operations use address values. The index value is initially an arbitrary code tree but we can recognise common cases of constant index values or where a constant has been added to the index. TODO: If these are C memory moves we can also look at the base address. The base address for C memory operations is a LargeWord.word value i.e. the address is contained in a box. The base addresses for ML memory moves is an ML address i.e. unboxed. *) and simpAddress({base, index=NONE, offset}, _, context) = let val (genBase, decBase, _ (*specBase*)) = simpSpecial(base, context, RevList[]) in ({base=genBase, index=NONE, offset=offset}, decBase) end | simpAddress({base, index=SOME index, offset}, multiplier, context) = let val (genBase, RevList decBase, _) = simpSpecial(base, context, RevList[]) val (genIndex, RevList decIndex, _ (* specIndex *)) = simpSpecial(index, context, RevList[]) val (newIndex, newOffset) = case genIndex of Constnt(indexOffset, _) => if isShort indexOffset then (NONE, offset + toShort indexOffset * multiplier) else (SOME genIndex, offset) | _ => (SOME genIndex, offset) in ({base=genBase, index=newIndex, offset=newOffset}, RevList(decIndex @ decBase)) end (* (* A built-in function. We can call certain built-ins immediately if the arguments are constants. *) and simpBuiltIn(rtsCallNo, argList, context as { reprocess, ...}) = let val copiedArgs = map (fn arg => simpSpecial(arg, context)) argList open RuntimeCalls (* When checking for a constant we need to check that there are no bindings. They could have side-effects. *) fun isAConstant(Constnt _, [], _) = true | isAConstant _ = false in (* If the function is an RTS call that is safe to evaluate immediately and all the arguments are constants evaluate it now. *) if earlyRtsCall rtsCallNo andalso List.all isAConstant copiedArgs then let val () = reprocess := true exception Interrupt = Thread.Thread.Interrupt (* Turn the arguments into a vector. *) val argVector = case makeConstVal(mkTuple(List.map specialToGeneral copiedArgs)) of Constnt(w, _) => w | _ => raise InternalError "makeConstVal: Not constant" (* Call the function. If it raises an exception (e.g. divide by zero) generate code to raise the exception at run-time. We don't do that for Interrupt which we assume only arises by user interaction and not as a result of executing the code so we reraise that exception immediately. *) val ioOp : int -> machineWord = RunCall.run_call1 RuntimeCalls.POLY_SYS_io_operation (* We need callcode_tupled here because we pass the arguments as a tuple but the RTS functions we're calling expect arguments in registers or on the stack. *) val call: (address * machineWord) -> machineWord = RunCall.run_call1 RuntimeCalls.POLY_SYS_callcode_tupled val code = Constnt (call(toAddress(ioOp rtsCallNo), argVector), []) handle exn as Interrupt => raise exn (* Must not handle this *) | exn => Raise (Constnt(toMachineWord exn, [])) in (code, [], EnvSpecNone) end (* We can optimise certain built-ins in combination with others. If we have POLY_SYS_unsigned_to_longword combined with POLY_SYS_longword_to_tagged we can eliminate both. This can occur in cases such as Word.fromLargeWord o Word8.toLargeWord. If we have POLY_SYS_cmem_load_X functions where the address is formed by adding a constant to an address we can move the addend into the load instruction. *) (* TODO: Could we also have POLY_SYS_signed_to_longword here? *) else if rtsCallNo = POLY_SYS_longword_to_tagged andalso (case copiedArgs of [(_, _, EnvSpecBuiltIn(r, _))] => r = POLY_SYS_unsigned_to_longword | _ => false) then let val arg = (* Get the argument of the argument. *) case copiedArgs of [(_, _, EnvSpecBuiltIn(_, [arg]))] => arg | _ => raise Bind in (arg, [], EnvSpecNone) end else if (rtsCallNo = POLY_SYS_cmem_load_8 orelse rtsCallNo = POLY_SYS_cmem_load_16 orelse rtsCallNo = POLY_SYS_cmem_load_32 orelse rtsCallNo = POLY_SYS_cmem_load_64 orelse rtsCallNo = POLY_SYS_cmem_store_8 orelse rtsCallNo = POLY_SYS_cmem_store_16 orelse rtsCallNo = POLY_SYS_cmem_store_32 orelse rtsCallNo = POLY_SYS_cmem_store_64) andalso (* Check if the first argument is an addition. The second should be a constant. If the addend is a constant it will be a large integer i.e. the address of a byte segment. *) let (* Check that we have a valid value to add to a large word. The cmem_load/store values sign extend their arguments so we use toLargeWordX here. *) fun isAcceptableOffset c = if isShort c (* Shouldn't occur. *) then false else let val l: LargeWord.word = RunCall.unsafeCast c in Word.toLargeWordX(Word.fromLargeWord l) = l end in case copiedArgs of (_, _, EnvSpecBuiltIn(r, args)) :: (Constnt _, _, _) :: _ => r = POLY_SYS_plus_longword andalso (case args of (* If they were both constants we'd have folded them. *) [Constnt(c, _), _] => isAcceptableOffset c | [_, Constnt(c, _)] => isAcceptableOffset c | _ => false) | _ => false end then let (* We have a load or store with an added constant. *) val (base, offset) = case copiedArgs of (_, _, EnvSpecBuiltIn(_, [Constnt(offset, _), base])) :: (Constnt(existing, _), _, _) :: _ => (base, Word.fromLargeWord(RunCall.unsafeCast offset) + toShort existing) | (_, _, EnvSpecBuiltIn(_, [base, Constnt(offset, _)])) :: (Constnt(existing, _), _, _) :: _ => (base, Word.fromLargeWord(RunCall.unsafeCast offset) + toShort existing) | _ => raise Bind val newDecs = List.map(fn h => makeNewDecl(h, context)) copiedArgs val genArgs = List.map(fn ((g, _), _) => envGeneralToCodetree g) newDecs val preDecs = List.foldr (op @) [] (List.map #2 newDecs) val gen = BuiltIn(rtsCallNo, base :: Constnt(toMachineWord offset, []) :: List.drop(genArgs, 2)) in (gen, preDecs, EnvSpecNone) end else let (* Create bindings for the arguments. This ensures that any side-effects in the evaluation of the arguments are performed in the correct order even if the application of the built-in itself is applicative. The new arguments are either loads or constants which are applicative. *) val newDecs = List.map(fn h => makeNewDecl(h, context)) copiedArgs val genArgs = List.map(fn ((g, _), _) => envGeneralToCodetree g) newDecs val preDecs = List.foldr (op @) [] (List.map #2 newDecs) val gen = BuiltIn(rtsCallNo, genArgs) val spec = if reorderable gen then EnvSpecBuiltIn(rtsCallNo, genArgs) else EnvSpecNone in (gen, preDecs, spec) end end *) and simpIfThenElse(condTest, condThen, condElse, context, tailDecs) = (* If-then-else. The main simplification is if we have constants in the test or in both the arms. *) let val word0 = toMachineWord 0 val word1 = toMachineWord 1 val False = word0 val True = word1 in case simpSpecial(condTest, context, tailDecs) of (* If the test is a constant we can return the appropriate arm and ignore the other. *) (Constnt(testResult, _), bindings, _) => let val arm = if wordEq (testResult, False) (* false - return else-part *) then condElse (* if false then x else y == y *) (* if true then x else y == x *) else condThen in simpSpecial(arm, context, bindings) end | (testGen, testbindings as RevList testBList, testSpec) => let fun mkNot (Unary{oper=BuiltIns.NotBoolean, arg1}) = arg1 | mkNot arg = Unary{oper=BuiltIns.NotBoolean, arg1=arg} (* If the test involves a variable that was created with a NOT it's better to move it in here. *) val testCond = case testSpec of EnvSpecUnary(BuiltIns.NotBoolean, arg1) => mkNot arg1 | _ => testGen in case (simpSpecial(condThen, context, RevList[]), simpSpecial(condElse, context, RevList[])) of ((thenConst as Constnt(thenVal, _), RevList [], _), (elseConst as Constnt(elseVal, _), RevList [], _)) => (* Both arms return constants. This situation can arise in situations where we have andalso/orelse where the second "argument" has been reduced to a constant. *) if wordEq (thenVal, elseVal) then (* If the test has a side-effect we have to do it otherwise we can remove it. If we're in a nested andalso/orelse that may mean we can simplify the next level out. *) (thenConst (* or elseConst *), if sideEffectFree testCond then testbindings else RevList(NullBinding testCond :: testBList), EnvSpecNone) (* if x then true else false == x *) else if wordEq (thenVal, True) andalso wordEq (elseVal, False) then (testCond, testbindings, EnvSpecNone) (* if x then false else true == not x *) else if wordEq (thenVal, False) andalso wordEq (elseVal, True) then (mkNot testCond, testbindings, EnvSpecNone) else (* can't optimise *) (Cond (testCond, thenConst, elseConst), testbindings, EnvSpecNone) (* Rewrite "if x then raise y else z" into "(if x then raise y else (); z)" The advantage is that any tuples in z are lifted outside the "if". *) | (thenPart as (Raise _, _:revlist, _), (elsePart, RevList elseBindings, elseSpec)) => (* then-part raises an exception *) (elsePart, RevList(elseBindings @ NullBinding(Cond (testCond, specialToGeneral thenPart, CodeZero)) :: testBList), elseSpec) | ((thenPart, RevList thenBindings, thenSpec), elsePart as (Raise _, _, _)) => (* else part raises an exception *) (thenPart, RevList(thenBindings @ NullBinding(Cond (testCond, CodeZero, specialToGeneral elsePart)) :: testBList), thenSpec) | (thenPart, elsePart) => (Cond (testCond, specialToGeneral thenPart, specialToGeneral elsePart), testbindings, EnvSpecNone) end end (* Tuple construction. Tuples are also used for datatypes and structures (i.e. modules) *) and simpTuple(entries, isVariant, context, tailDecs) = (* The main reason for optimising record constructions is that they appear as tuples in ML. We try to ensure that loads from locally created tuples do not involve indirecting from the tuple but can get the value which was put into the tuple directly. If that is successful we may find that the tuple is never used directly so the use-count mechanism will ensure it is never created. *) let val tupleSize = List.length entries (* The record construction is treated as a block of local declarations so that any expressions which might have side-effects are done exactly once. *) (* We thread the bindings through here to avoid having to append the result. *) fun processFields([], bindings) = ([], bindings) | processFields(field::fields, bindings) = let val (thisField, newBindings) = makeNewDecl(simpSpecial(field, context, bindings), context) val (otherFields, resBindings) = processFields(fields, newBindings) in (thisField::otherFields, resBindings) end val (fieldEntries, allBindings) = processFields(entries, tailDecs) (* Make sure we include any inline code in the result. If this tuple is being "exported" we will lose the "special" part. *) fun envResToCodetree(EnvGenLoad(ext), _) = Extract ext | envResToCodetree(EnvGenConst(w, p), s) = Constnt(w, setInline s p) val generalFields = List.map envResToCodetree fieldEntries val genRec = if List.all isConstnt generalFields then makeConstVal(Tuple{ fields = generalFields, isVariant = isVariant }) else Tuple{ fields = generalFields, isVariant = isVariant } (* Get the field from the tuple if possible. If it's a variant, though, we may try to get an invalid field. See Tests/Succeed/Test167. *) fun getField addr = if addr < tupleSize then List.nth(fieldEntries, addr) else if isVariant then (EnvGenConst(toMachineWord 0, []), EnvSpecNone) else raise InternalError "getField - invalid index" val specRec = EnvSpecTuple(tupleSize, getField) in (genRec, allBindings, specRec) end and simpFieldSelect(base, offset, indKind, context, tailDecs) = let val (genSource, decSource, specSource) = simpSpecial(base, context, tailDecs) in (* Try to do the selection now if possible. *) case specSource of EnvSpecTuple(_, recEnv) => let (* The "special" entry we've found is a tuple. That means that we are taking a field from a tuple we made earlier and so we should be able to get the original code we used when we made the tuple. That might mean the tuple is never used and we can optimise away the construction of it completely. *) val (newGen, newSpec) = recEnv offset in (envGeneralToCodetree newGen, decSource, newSpec) end | _ => (* No special case possible. If the tuple is a constant mkInd/mkVarField will do the selection immediately. *) let val genSelect = case indKind of IndTuple => mkInd(offset, genSource) | IndVariant => mkVarField(offset, genSource) | IndContainer => mkIndContainer(offset, genSource) in (genSelect, decSource, EnvSpecNone) end end (* Process a SetContainer. Unlike the other simpXXX functions this is called after the arguments have been processed. We try to push the SetContainer to the leaves of the expression. This is particularly important with tail-recursive functions that return tuples. Without this the function will lose tail-recursion since each recursion will be followed by code to copy the result back to the previous container. *) and simpPostSetContainer(container, Tuple{fields, ...}, RevList tupleDecs, filter) = let (* Apply the filter now. *) fun select(n, hd::tl) = if n >= BoolVector.length filter then [] else if BoolVector.sub(filter, n) then hd :: select(n+1, tl) else select(n+1, tl) | select(_, []) = [] val selected = select(0, fields) (* Frequently we will have produced an indirection from the same base. These will all be bindings so we have to reverse the process. *) fun findOriginal a = List.find(fn Declar{addr, ...} => addr = a | _ => false) tupleDecs fun checkFields(last, Extract(LoadLocal a) :: tl) = ( case findOriginal a of SOME(Declar{value=Indirect{base=Extract ext, indKind=IndContainer, offset, ...}, ...}) => ( case last of NONE => checkFields(SOME(ext, [offset]), tl) | SOME(lastExt, offsets) => (* It has to be the same base and with increasing offsets (no reordering). *) if lastExt = ext andalso offset > hd offsets then checkFields(SOME(ext, offset :: offsets), tl) else NONE ) | _ => NONE ) | checkFields(_, _ :: _) = NONE | checkFields(last, []) = last fun fieldsToFilter fields = let val maxDest = List.foldl Int.max ~1 fields val filterArray = BoolArray.array(maxDest+1, false) val _ = List.app(fn n => BoolArray.update(filterArray, n, true)) fields in BoolArray.vector filterArray end in case checkFields(NONE, selected) of SOME (ext, fields) => (* It may be a container. *) let val filter = fieldsToFilter fields in case ext of LoadLocal localAddr => let (* Is this a container? If it is and we're copying all of it we can replace the inner container with a binding to the outer. We have to be careful because it is possible that we may create and set the inner container, then have some bindings that do some side-effects with the inner container before then copying it to the outer container. For simplicity and to maintain the condition that the container is set in the tails we only merge the containers if it's at the end (after any "filtering"). *) val allSet = BoolVector.foldl (fn (a, t) => a andalso t) true filter fun findContainer [] = NONE | findContainer (Declar{value, ...} :: tl) = if sideEffectFree value then findContainer tl else NONE | findContainer (Container{addr, size, setter, ...} :: tl) = if localAddr = addr andalso size = BoolVector.length filter andalso allSet then SOME (setter, tl) else NONE | findContainer _ = NONE in case findContainer tupleDecs of SOME (setter, decs) => (* Put in a binding for the inner container address so the setter will set the outer container. For this to work all loads from the stack must use native word length. *) mkEnv(List.rev(Declar{addr=localAddr, value=container, use=[]} :: decs), setter) | NONE => mkEnv(List.rev tupleDecs, SetContainer{container=container, tuple = mkTuple selected, filter=BoolVector.tabulate(List.length selected, fn _ => true)}) end | _ => mkEnv(List.rev tupleDecs, SetContainer{container=container, tuple = mkTuple selected, filter=BoolVector.tabulate(List.length selected, fn _ => true)}) end | NONE => mkEnv(List.rev tupleDecs, SetContainer{container=container, tuple = mkTuple selected, filter=BoolVector.tabulate(List.length selected, fn _ => true)}) end | simpPostSetContainer(container, Cond(ifpt, thenpt, elsept), RevList tupleDecs, filter) = mkEnv(List.rev tupleDecs, Cond(ifpt, simpPostSetContainer(container, thenpt, RevList [], filter), simpPostSetContainer(container, elsept, RevList [], filter))) | simpPostSetContainer(container, Newenv(envDecs, envExp), RevList tupleDecs, filter) = simpPostSetContainer(container, envExp, RevList(List.rev envDecs @ tupleDecs), filter) | simpPostSetContainer(container, BeginLoop{loop, arguments}, RevList tupleDecs, filter) = mkEnv(List.rev tupleDecs, BeginLoop{loop = simpPostSetContainer(container, loop, RevList [], filter), arguments=arguments}) | simpPostSetContainer(_, loop as Loop _, RevList tupleDecs, _) = (* If we are inside a BeginLoop we only set the container on leaves that exit the loop. Loop entries will go back to the BeginLoop so we don't add SetContainer nodes. *) mkEnv(List.rev tupleDecs, loop) | simpPostSetContainer(container, Handle{exp, handler, exPacketAddr}, RevList tupleDecs, filter) = mkEnv(List.rev tupleDecs, Handle{ exp = simpPostSetContainer(container, exp, RevList [], filter), handler = simpPostSetContainer(container, handler, RevList [], filter), exPacketAddr = exPacketAddr}) | simpPostSetContainer(container, tupleGen, RevList tupleDecs, filter) = mkEnv(List.rev tupleDecs, mkSetContainer(container, tupleGen, filter)) fun simplifier(c, numLocals) = let val localAddressAllocator = ref 0 val addrTab = Array.array(numLocals, NONE) fun lookupAddr (LoadLocal addr) = valOf(Array.sub(addrTab, addr)) | lookupAddr (env as LoadArgument _) = (EnvGenLoad env, EnvSpecNone) | lookupAddr (env as LoadRecursive) = (EnvGenLoad env, EnvSpecNone) | lookupAddr (LoadClosure _) = raise InternalError "top level reached in simplifier" and enterAddr (addr, tab) = Array.update (addrTab, addr, SOME tab) fun mkAddr () = ! localAddressAllocator before localAddressAllocator := ! localAddressAllocator + 1 val reprocess = ref false val (gen, RevList bindings, spec) = simpSpecial(c, {lookupAddr = lookupAddr, enterAddr = enterAddr, nextAddress = mkAddr, reprocess = reprocess}, RevList[]) in ((gen, List.rev bindings, spec), ! localAddressAllocator, !reprocess) end fun specialToGeneral(g, b as _ :: _, s) = mkEnv(b, specialToGeneral(g, [], s)) | specialToGeneral(Constnt(w, p), [], s) = Constnt(w, setInline s p) | specialToGeneral(g, [], _) = g structure Sharing = struct type codetree = codetree and codeBinding = codeBinding and envSpecial = envSpecial end end; diff --git a/mlsource/MLCompiler/CodeTree/CODETREE_STATIC_LINK_AND_CASES.sml b/mlsource/MLCompiler/CodeTree/CODETREE_STATIC_LINK_AND_CASES.sml index 4f38c18d..d786e252 100644 --- a/mlsource/MLCompiler/CodeTree/CODETREE_STATIC_LINK_AND_CASES.sml +++ b/mlsource/MLCompiler/CodeTree/CODETREE_STATIC_LINK_AND_CASES.sml @@ -1,879 +1,884 @@ (* Copyright (c) 2012-13, 2015-17 David C.J. Matthews This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) functor CODETREE_STATIC_LINK_AND_CASES( structure BASECODETREE: BaseCodeTreeSig structure CODETREE_FUNCTIONS: CodetreeFunctionsSig structure GCODE: GENCODESIG structure DEBUG: DEBUGSIG structure PRETTY : PRETTYSIG structure BACKENDTREE: BackendIntermediateCodeSig sharing BASECODETREE.Sharing = CODETREE_FUNCTIONS.Sharing = GCODE.Sharing = PRETTY.Sharing = BACKENDTREE.Sharing ) : CodegenTreeSig = struct open BASECODETREE open Address open BACKENDTREE datatype caseType = datatype BACKENDTREE.caseType exception InternalError = Misc.InternalError open BACKENDTREE.CodeTags (* Property tag to indicate which arguments to a function are functions that are only ever called. *) val closureFreeArgsTag: int list Universal.tag = Universal.tag() datatype maybeCase = IsACase of { cases : (backendIC * word) list, test : backendIC, caseType: caseType, default : backendIC } | NotACase of backendIC fun staticLinkAndCases (pt, localAddressCount) = let fun copyCode (pt, nonLocals, recursive, localCount, localAddresses, argClosure) = let (* "closuresForLocals" is a flag indicating that if the declaration is a function a closure must be made for it. *) val closuresForLocals = Array.array(localCount, false) val newLocalAddresses = Array.array (localCount, 0) val argProperties = Array.array(localCount, []) (* Reference to local or non-local bindings. This sets the "closure" property on the binding depending on how the binding will be used. *) fun locaddr (LoadLocal addr, closure) = let val () = if closure then Array.update (closuresForLocals, addr, true) else () val newAddr = Array.sub(newLocalAddresses, addr) in BICLoadLocal newAddr end | locaddr(LoadArgument addr, closure) = ( argClosure(addr, closure); BICLoadArgument addr ) | locaddr(LoadRecursive, closure) = recursive closure | locaddr(LoadClosure addr, closure) = #1 (nonLocals (addr, closure)) (* Argument properties. This returns information of which arguments can have functions passed in without requiring a full heap closure. *) fun argumentProps(LoadLocal addr) = Array.sub(argProperties, addr) | argumentProps(LoadArgument _) = [] | argumentProps LoadRecursive = [] | argumentProps (LoadClosure addr) = #2 (nonLocals (addr, false)) fun makeDecl addr = let val newAddr = ! localAddresses before (localAddresses := !localAddresses+1) val () = Array.update (closuresForLocals, addr, false) val () = Array.update (newLocalAddresses, addr, newAddr) val () = Array.update (argProperties, addr, []) in newAddr end fun insert(Eval { function = Extract LoadRecursive, argList, resultType, ...}) = let (* Recursive. If we pass an argument in the same position we don't necessarily need a closure. It depends on what else happens to it. *) fun mapArgs(n, (Extract (ext as LoadArgument m), t) :: tail) = (BICExtract(locaddr(ext, n <> m)), t) :: mapArgs(n+1, tail) | mapArgs(n, (c, t) :: tail) = (insert c, t) :: mapArgs(n+1, tail) | mapArgs(_, []) = [] val newargs = mapArgs(0, argList) val func = locaddr(LoadRecursive, (* closure = *) false) in (* If we are calling a function which has been declared this does not require it to have a closure. Any other use of the function would. *) BICEval {function = BICExtract func, argList = newargs, resultType=resultType} end | insert(Eval { function = Extract ext, argList, resultType, ...}) = let (* Non-recursive but a binding. *) val cfArgs = argumentProps ext fun isIn n = not(List.exists(fn m => m = n) cfArgs) fun mapArgs(n, (Extract ext, t) :: tail) = (BICExtract(locaddr(ext, isIn n)), t) :: mapArgs(n+1, tail) | mapArgs(n, (Lambda lam, t) :: tail) = (insertLambda(lam, isIn n), t) :: mapArgs(n+1, tail) | mapArgs(n, (c, t) :: tail) = (insert c, t) :: mapArgs(n+1, tail) | mapArgs(_, []) = [] val newargs = mapArgs(0, argList) val func = locaddr(ext, (* closure = *) false) in (* If we are calling a function which has been declared this does not require it to have a closure. Any other use of the function would. *) BICEval {function = BICExtract func, argList = newargs, resultType=resultType} end | insert(Eval { function = Constnt(w, p), argList, resultType, ...}) = let (* Constant function. *) val cfArgs = case List.find (Universal.tagIs closureFreeArgsTag) p of NONE => [] | SOME u => Universal.tagProject closureFreeArgsTag u fun isIn n = not(List.exists(fn m => m = n) cfArgs) fun mapArgs(n, (Extract ext, t) :: tail) = (BICExtract(locaddr(ext, isIn n)), t) :: mapArgs(n+1, tail) | mapArgs(n, (Lambda lam, t) :: tail) = (insertLambda(lam, isIn n), t) :: mapArgs(n+1, tail) | mapArgs(n, (c, t) :: tail) = (insert c, t) :: mapArgs(n+1, tail) | mapArgs(_, []) = [] val newargs = mapArgs(0, argList) in BICEval {function = BICConstnt (w, p), argList = newargs, resultType=resultType} end | insert(Eval { function = Lambda lam, argList, resultType, ...}) = let (* Call of a lambda. Typically this will be a recursive function that can't be inlined. *) val newargs = map(fn (c, t) => (insert c, t)) argList val (copiedLambda, newClosure, makeRecClosure, _) = copyLambda lam val func = copyProcClosure (copiedLambda, newClosure, makeRecClosure) in BICEval {function = func, argList = newargs, resultType=resultType} end | insert(Eval { function, argList, resultType, ...}) = let (* Process the arguments first. *) val newargs = map(fn (c, t) => (insert c, t)) argList val func = insert function in BICEval {function = func, argList = newargs, resultType=resultType} end | insert GetThreadId = BICGetThreadId | insert(Unary { oper, arg1 }) = BICUnary { oper = oper, arg1 = insert arg1 } | insert(Binary { oper, arg1, arg2 }) = BICBinary { oper = oper, arg1 = insert arg1, arg2 = insert arg2 } | insert(Arbitrary { oper=ArbCompare test, shortCond, arg1, arg2, longCall}) = let val insArg1 = insert arg1 and insArg2 = insert arg2 and insCall = insert longCall and insShort = insert shortCond (* We have to rewrite this. *) (* if isShort i andalso isShort j then toShort i < toShort j else callComp(i, j) < 0 *) fun fixedComp(arg1, arg2) = BICBinary { oper = BuiltIns.WordComparison{test=test, isSigned=true}, arg1 = arg1, arg2 = arg2 } val zeroFalse = BICConstnt(toMachineWord 0, []) in BICCond( insShort, fixedComp(insArg1, insArg2), fixedComp(insCall, zeroFalse) ) end | insert(Arbitrary { oper=ArbArith arith, shortCond, arg1, arg2, longCall}) = let val insArg1 = insert arg1 and insArg2 = insert arg2 and insCall = insert longCall and insShort = insert shortCond in BICArbitrary{oper=arith, shortCond=insShort, arg1=insArg1, arg2=insArg2, longCall=insCall} end | insert(AllocateWordMemory {numWords, flags, initial}) = BICAllocateWordMemory { numWords = insert numWords, flags = insert flags, initial = insert initial } | insert(Extract ext) = (* Load the value bound to an identifier. The closure flag is set to true since the only cases where a closure is not needed, eval and load-andStore, are handled separately. *) BICExtract(locaddr(ext, (* closure = *) true)) | insert(Indirect {base, offset, indKind=IndContainer}) = BICLoadContainer {base = insert base, offset = offset} | insert(Indirect {base, offset, ...}) = BICField {base = insert base, offset = offset} | insert(Constnt wp) = BICConstnt wp (* Constants can be returned untouched. *) | insert(BeginLoop{loop=body, arguments=argList, ...}) = (* Start of tail-recursive inline function. *) let (* Make entries in the tables for the arguments. *) val newAddrs = List.map (fn ({addr, ...}, _) => makeDecl addr) argList (* Process the body. *) val insBody = insert body (* Finally the initial argument values. *) local fun copyDec(({value, ...}, t), addr) = ({addr=addr, value=insert value}, t) in val newargs = ListPair.map copyDec (argList, newAddrs) end in (* Add the kill entries on after the loop. *) BICBeginLoop{loop=insBody, arguments=newargs} end | insert(Loop argList) = (* Jump back to start of tail-recursive function. *) BICLoop(List.map(fn (c, t) => (insert c, t)) argList) | insert(Raise x) = BICRaise (insert x) (* See if we can use a case-instruction. Arguably this belongs in the optimiser but it is only really possible when we have removed redundant declarations. *) | insert(Cond(condTest, condThen, condElse)) = reconvertCase(copyCond (condTest, condThen, condElse)) | insert(Newenv(ptElist, ptExp)) = let (* Process the body. Recurses down the list of declarations and expressions processing each, and then reconstructs the list on the way back. *) fun copyDeclarations ([]) = [] | copyDeclarations (Declar({addr=caddr, value = Lambda lam, ...}) :: vs) = let (* Binding a Lambda - process the function first. *) val newAddr = makeDecl caddr val (copiedLambda, newClosure, makeRecClosure, cfArgs) = copyLambda lam val () = Array.update(argProperties, caddr, cfArgs) (* Process all the references to the function. *) val rest = copyDeclarations vs (* We now know if we need a heap closure. *) val dec = copyProcClosure(copiedLambda, newClosure, makeRecClosure orelse Array.sub(closuresForLocals, caddr)) in BICDeclar{addr=newAddr, value=dec} :: rest end | copyDeclarations (Declar({addr=caddr, value = pt, ...}) :: vs) = let (* Non-function binding. *) val newAddr = makeDecl caddr val rest = copyDeclarations vs in BICDeclar{addr=newAddr, value=insert pt} :: rest end | copyDeclarations (RecDecs mutualDecs :: vs) = let (* Mutually recursive declarations. Any of the declarations may refer to any of the others. This causes several problems in working out the use-counts and whether the functions (they should be functions) need closures. A function will need a closure if any reference would require one (i.e. does anything other than call it). The reference may be from one of the other mutually recursive declarations and may be because that function requires a full closure. This means that once we have dealt with any references in the rest of the containing block we have to repeatedly scan the list of declarations removing those which need closures until we are left with those that do not. The use-counts can only be obtained when all the non-local lists have been copied. *) (* First go down the list making a declaration for each entry. This makes sure there is a table entry for all the declarations. *) val _ = List.map (fn {addr, ...} => makeDecl addr) mutualDecs (* Process the rest of the block. Identifies all other references to these declarations. *) val restOfBlock = copyDeclarations vs (* We now want to find out which of the declarations require closures. First we copy all the declarations, except that we don't copy the non-local lists of functions. *) fun copyDec ({addr=caddr, lambda, ...}) = let val (dec, newClosure, makeRecClosure, cfArgs) = copyLambda lambda val () = if makeRecClosure then Array.update (closuresForLocals, caddr, true) else () val () = Array.update(argProperties, caddr, cfArgs) in (caddr, dec, newClosure) end val copiedDecs = map copyDec mutualDecs (* We now have identified all possible references to the functions apart from those of the closures themselves. Any of closures may refer to any other function so we must iterate until all the functions which need full closures have been processed. *) fun processClosures([], outlist, true) = (* Sweep completed. - Must repeat. *) processClosures(outlist, [], false) | processClosures([], outlist, false) = (* We have processed the whole of the list without finding anything which needs a closure. The remainder do not need full closures. *) let fun mkLightClosure ((addr, value, newClosure)) = let val clos = copyProcClosure(value, newClosure, false) val newAddr = Array.sub(newLocalAddresses, addr) in {addr=newAddr, value=clos} end in map mkLightClosure outlist end | processClosures((h as (caddr, value, newClosure))::t, outlist, someFound) = if Array.sub(closuresForLocals, caddr) then let (* Must copy it. *) val clos = copyProcClosure(value, newClosure, true) val newAddr = Array.sub(newLocalAddresses, caddr) in {addr=newAddr, value=clos} :: processClosures(t, outlist, true) end (* Leave it for the moment. *) else processClosures(t, h :: outlist, someFound) val decs = processClosures(copiedDecs, [], false) local fun isLambda{value=BICLambda _, ...} = true | isLambda _ = false in val (lambdas, nonLambdas) = List.partition isLambda decs end fun asMutual{addr, value = BICLambda lambda} = {addr=addr, lambda=lambda} | asMutual _ = raise InternalError "asMutual" in (* Return the mutual declarations and the rest of the block. *) if null lambdas then map BICDeclar nonLambdas @ restOfBlock (* None left *) else BICRecDecs (map asMutual lambdas) :: (map BICDeclar nonLambdas @ restOfBlock) end (* copyDeclarations.isMutualDecs *) | copyDeclarations (NullBinding v :: vs) = let (* Not a declaration - process this and the rest. *) (* Must process later expressions before earlier ones so that the last references to variables are found correctly. DCJM 30/11/99. *) val copiedRest = copyDeclarations vs; val copiedNode = insert v in (* Expand out blocks *) case copiedNode of BICNewenv(decs, exp) => decs @ (BICNullBinding exp :: copiedRest) | _ => BICNullBinding copiedNode :: copiedRest end | copyDeclarations (Container{addr, size, setter, ...} :: vs) = let val newAddr = makeDecl addr val rest = copyDeclarations vs val setCode = insert setter in BICDecContainer{addr=newAddr, size=size} :: BICNullBinding setCode :: rest end val insElist = copyDeclarations(ptElist @ [NullBinding ptExp]) fun mkEnv([], exp) = exp | mkEnv(decs, exp) = BICNewenv(decs, exp) fun decSequenceWithFinalExp decs = let fun splitLast _ [] = raise InternalError "decSequenceWithFinalExp: empty" | splitLast decs [BICNullBinding exp] = (List.rev decs, exp) | splitLast _ [_] = raise InternalError "decSequenceWithFinalExp: last is not a NullDec" | splitLast decs (hd::tl) = splitLast (hd:: decs) tl in mkEnv(splitLast [] decs) end in (* TODO: Tidy this up. *) decSequenceWithFinalExp insElist end (* isNewEnv *) | insert(Tuple { fields, ...}) = BICTuple (map insert fields) | insert(Lambda lam) = (* Using a lambda in a context other than a call or being passed to a function that is known only to call the function. It requires a heap closure. *) insertLambda(lam, true) | insert(Handle { exp, handler, exPacketAddr }) = let (* The order here is important. We want to make sure that the last reference to a variable really is the last. *) val newAddr = makeDecl exPacketAddr val hand = insert handler val exp = insert exp in BICHandle {exp = exp, handler = hand, exPacketAddr=newAddr} end | insert(SetContainer {container, tuple, filter}) = BICSetContainer{container = insert container, tuple = insert tuple, filter = filter} | insert(TagTest{test, tag, maxTag}) = BICTagTest{test=insert test, tag=tag, maxTag=maxTag} | insert(LoadOperation{kind, address}) = BICLoadOperation{kind=kind, address=insertAddress address} | insert(StoreOperation{kind, address, value}) = BICStoreOperation{kind=kind, address=insertAddress address, value=insert value} | insert(BlockOperation{kind, sourceLeft, destRight, length}) = BICBlockOperation{ kind=kind, sourceLeft=insertAddress sourceLeft, destRight=insertAddress destRight, length=insert length} and insertLambda (lam, needsClosure) = let val (copiedLambda, newClosure, _, _) = copyLambda lam in copyProcClosure (copiedLambda, newClosure, needsClosure) end and insertAddress{base, index, offset} = {base=insert base, index=Option.map insert index, offset=offset} and copyCond (condTest, condThen, condElse): maybeCase = let (* Process the then-part. *) val insThen = insert condThen (* Process the else-part. If it's a conditional process it here. *) val insElse = case condElse of Cond(i, t, e) => copyCond(i, t, e) | _ => NotACase(insert condElse) (* Process the condition after the then- and else-parts. *) val insFirst = insert condTest type caseVal = { tag: word, test: codetree, caseType: caseType } option; (* True if both instructions are loads or indirections with the same effect. More complicated cases could be considered but function calls must always be treated as different. Note: the reason we consider Indirect entries here as well as Extract is because we (used to) defer Indirect entries. *) datatype similarity = Different | Similar of bicLoadForm fun similar (BICExtract a, BICExtract b) = if a = b then Similar a else Different | similar (BICField{offset=aOff, base=aBase}, BICField{offset=bOff, base=bBase}) = if aOff <> bOff then Different else similar (aBase, bBase) | similar _ = Different; (* If we have a call to the int equality operation then we may be able to use an indexed case. N.B. This works equally for word values (unsigned) and fixed precision int (unsigned) but is unsafe for arbitrary precision since - the lower levels assume that all values are tagged. *) + the lower levels assume that all values are tagged. + This could be used for PointerEq which is what arbitrary precision will generate + provided that there was an extra check for long values. N.B. the same also + happens for + e.g. datatype t = A | B | C | D | E of int*int + i.e. one non-nullary constructor. *) fun findCase (BICBinary{oper=BuiltIns.WordComparison{test=BuiltIns.TestEqual, ...}, arg1, arg2}) = let in case (arg1, arg2) of (BICConstnt(c1, _), arg2) => if isShort c1 then SOME{tag=toShort c1, test=arg2, caseType = CaseWord} else NONE (* Not a short constant. *) | (arg1, BICConstnt(c2, _)) => if isShort c2 then SOME{tag=toShort c2, test=arg1, caseType = CaseWord} else NONE (* Not a short constant. *) | _ => NONE (* Wrong number of arguments - should raise exception? *) end | findCase(BICTagTest { test, tag, maxTag }) = SOME { tag=tag, test=test, caseType=CaseTag maxTag } | findCase _ = NONE val testCase = findCase insFirst in case testCase of NONE => (* Can't use a case *) NotACase(BICCond (insFirst, insThen, reconvertCase insElse)) | SOME { tag=caseTags, test=caseTest, caseType=caseCaseTest } => (* Can use a case. Can we combine two cases? If we have an expression like "if x = a then .. else if x = b then ..." we can combine them into a single "case". *) case insElse of IsACase { cases=nextCases, test=nextTest, default=nextDefault, caseType=nextCaseType } => ( case (similar(nextTest, caseTest), caseCaseTest = nextCaseType) of (* Note - it is legal (though completely redundant) for the same case to appear more than once in the list. This is not checked for at this stage. *) (Similar _, true) => IsACase { cases = (insThen, caseTags) :: map (fn (c, l) => (c, l)) nextCases, test = nextTest, default = nextDefault, caseType = caseCaseTest } | _ => (* Two case expressions but they test different variables. We can't combine them. *) IsACase { cases = [(insThen, caseTags)], test = caseTest, default = reconvertCase insElse, caseType=caseCaseTest } ) | NotACase elsePart => (* insElse is not a case *) IsACase { cases = [(insThen, caseTags)], test = caseTest, default = elsePart, caseType=caseCaseTest } end (* Check something that's been created as a Case and see whether it is sparse. If it is turn it back into a sequence of conditionals. This was previously done at the bottom level and the choice of when to use an indexed case was made by the architecture-specific code-generator. That's probably unnecessary and complicates the code-generator. *) and reconvertCase(IsACase{cases, test, default, caseType}) = let (* Count the number of cases and compute the maximum and minimum. *) (* If we are testing on integers we could have negative values here. Because we're using "word" here any negative values are treated as large positive values and so we won't use a "case". If this is a case on constructor tags we know the range. There will always be a "default" which may be anywhere in the range but if we construct a jump table that covers all the values we don't need the range checks. *) val useIndexedCase = case caseType of CaseTag _ => (* Exhaustive *) List.length cases > 4 | _ => let val (_, aLabel) = hd cases fun foldCases((_, w), (min, max)) = (Word.min(w, min), Word.max(w, max)) val (min, max) = List.foldl foldCases (aLabel, aLabel) cases val numberOfCases = List.length cases in numberOfCases > 7 andalso Word.fromInt numberOfCases >= (max - min) div 0w3 end in if useIndexedCase then let (* Create a contiguous range of labels. Eliminate any duplicates which are legal but redundant. *) local val labelCount = List.length cases (* Add an extra field before sorting which retains the ordering for equal labels. *) val ordered = ListPair.zipEq (cases, List.tabulate(labelCount, fn n=>n)) fun leq ((_, w1: word), n1: int) ((_, w2), n2) = if w1 = w2 then n1 <= n2 else w1 < w2 val sorted = List.map #1 (Misc.quickSort leq ordered) (* Filter out any duplicates. *) fun filter [] = [] | filter [p] = [p] | filter ((p as (_, lab1)) :: (q as (_, lab2)) :: tl) = if lab1 = lab2 then p :: filter tl else p :: filter (q :: tl) in val cases = filter sorted end val (isExhaustive, min, max) = case caseType of CaseTag max => (true, 0w0, max) | _ => let val (_, aLabel) = hd cases fun foldCases((_, w), (min, max)) = (Word.min(w, min), Word.max(w, max)) val (min, max) = List.foldl foldCases (aLabel, aLabel) cases in (false, min, max) end (* Create labels for each of the cases. Fill in any gaps with entries that will point to the default. We have to be careful if max happens to be the largest value of Word.word. In that case adding one to the range will give us a value less than max. *) fun extendCase(indexVal, cl as ((c, caseValue) :: cps)) = if indexVal + min = caseValue then SOME c :: extendCase(indexVal+0w1, cps) else NONE :: extendCase(indexVal+0w1, cl) | extendCase(indexVal, []) = (* We may not be at the end if this came from a CaseTag *) if indexVal > max-min then [] else NONE :: extendCase(indexVal+0w1, []) val fullCaseRange = extendCase(0w0, cases) val _ = Word.fromInt(List.length fullCaseRange) = max-min+0w1 orelse raise InternalError "Cases" in BICCase{cases=fullCaseRange, test=test, default=default, isExhaustive=isExhaustive, firstIndex=min} end else let fun reconvert [] = default | reconvert ((c, t) :: rest) = let val test = case caseType of CaseWord => BICBinary{ oper=BuiltIns.WordComparison{test=BuiltIns.TestEqual, isSigned=false}, arg1=test, arg2=BICConstnt(toMachineWord t, [])} | CaseTag maxTag => BICTagTest { test=test, tag=t, maxTag=maxTag } in BICCond(test, c, reconvert rest) end in reconvert cases end end | reconvertCase (NotACase t) = t (* Just a simple conditional. *) (* If "makeClosure" is true the function will need a full closure. It may need a full closure even if makeClosure is false if it involves a recursive reference which will need a closure. *) and copyLambda ({body=lambdaBody, argTypes, name=lambdaName, resultType, localCount, closure=lambdaClosure, ...}: lambdaForm) = let val newGrefs: loadForm list ref = ref [] (* non-local references *) val newNorefs = ref 0 (* number of non-local refs *) val makeClosureForRecursion = ref false (* A new table for the new function. *) fun prev (closureAddr, closure) = let val loadEntry = List.nth(lambdaClosure, closureAddr) (* Returns the closure address of the non-local *) fun makeClosureEntry([], _) = (* not found - construct new entry *) let val () = newGrefs := loadEntry :: !newGrefs; val newAddr = !newNorefs + 1; in newNorefs := newAddr; (* increment count *) newAddr-1 end | makeClosureEntry(oldEntry :: t, newAddr) = if oldEntry = loadEntry then newAddr-1 else makeClosureEntry(t, newAddr - 1) (* Set the closure flag if necessary and get the argument props. At this point we discard the "Load" entry returned by nonLocals and "recursive". The closure will be processed later. *) val argProps = case loadEntry of LoadLocal addr => let val () = if closure then Array.update (closuresForLocals, addr, true) else () in Array.sub(argProperties, addr) end | LoadArgument addr => (argClosure(addr, closure); []) | LoadRecursive => (recursive closure; []) | LoadClosure entry => #2 (nonLocals (entry, closure)) in (* Just return the closure entry. *) (BICLoadClosure(makeClosureEntry (!newGrefs, !newNorefs)), argProps) end fun recCall closure = (* Reference to the closure itself. *) ( if closure then makeClosureForRecursion := true else (); BICLoadRecursive ) local datatype tri = TriUnref | TriCall | TriClosure val argClosureArray = Array.array(List.length argTypes, TriUnref) in fun argClosure(n, t) = Array.update(argClosureArray, n, (* If this is true it requires a closure. If it is false it requires a closure if any other reference does. *) if t orelse Array.sub(argClosureArray, n) = TriClosure then TriClosure else TriCall) fun closureFreeArgs() = Array.foldri(fn (n, TriCall, l) => n :: l | (_, _, l) => l) [] argClosureArray end (* process the body *) val newLocalAddresses = ref 0 val (insertedCode, _) = copyCode (lambdaBody, prev, recCall, localCount, newLocalAddresses, argClosure) val globalRefs = !newGrefs val cfArgs = closureFreeArgs() in (BICLambda { body = insertedCode, name = lambdaName, closure = [], argTypes = map #1 argTypes, resultType = resultType, localCount = ! newLocalAddresses, heapClosure = false }, globalRefs, ! makeClosureForRecursion, cfArgs) end (* copyLambda *) (* Copy the closure of a function which has previously been processed by copyLambda. *) and copyProcClosure (BICLambda{ body, name, argTypes, resultType, localCount, ...}, newClosure, heapClosure) = let (* process the non-locals in this function *) (* If a heap closure is needed then any functions referred to from the closure also need heap closures.*) fun makeLoads ext = locaddr(ext, heapClosure) val copyRefs = rev (map makeLoads newClosure) in BICLambda { body = body, name = name, closure = copyRefs, argTypes = argTypes, resultType = resultType, localCount = localCount, heapClosure = heapClosure orelse null copyRefs (* False if closure is empty *) } end | copyProcClosure(pt, _, _) = pt (* may now be a constant *) (* end copyProcClosure *) in case pt of Lambda lam => let val (copiedLambda, newClosure, _, cfArgs) = copyLambda lam val code = copyProcClosure (copiedLambda, newClosure, true) val props = if null cfArgs then [] else [Universal.tagInject closureFreeArgsTag cfArgs] in (code, props) end | c as Newenv(_, exp) => let val code = insert c fun getProps(Extract(LoadLocal addr)) = let val cfArgs = Array.sub(argProperties, addr) in if null cfArgs then [] else [Universal.tagInject closureFreeArgsTag cfArgs] end | getProps(Tuple { fields, ...}) = let val fieldProps = map getProps fields in if List.all null fieldProps then [] else [Universal.tagInject CodeTags.tupleTag fieldProps] end | getProps _ = [] val props = getProps exp in (code, props) end | c as Constnt(_, p) => (insert c, p) | pt => (insert pt, []) end (* copyCode *) val outputAddresses = ref 0 fun topLevel _ = raise InternalError "outer level reached in copyCode" val (insertedCode, argProperties) = copyCode (pt, topLevel, topLevel, localAddressCount, outputAddresses, fn _ => ()) in (insertedCode, argProperties) end (* staticLinkAndCases *) type closureRef = GCODE.closureRef fun codeGenerate(lambda: lambdaForm, debugSwitches, closure) = let val (code, argProperties) = staticLinkAndCases(Lambda lambda, 0) val backendCode = code val () = if DEBUG.getParameter DEBUG.codetreeAfterOptTag debugSwitches then PRETTY.getCompilerOutput debugSwitches (BACKENDTREE.pretty backendCode) else () val bicLambda = case backendCode of BACKENDTREE.BICLambda lam => lam | _ => raise InternalError "Not BICLambda" val () = GCODE.gencodeLambda(bicLambda, debugSwitches, closure) in argProperties end structure Foreign = GCODE.Foreign (* Sharing can be copied from CODETREE. *) structure Sharing = struct open BASECODETREE.Sharing type closureRef = closureRef end end; diff --git a/mlsource/MLCompiler/CodeTree/X86Code/X86CodetreeToICode.ML b/mlsource/MLCompiler/CodeTree/X86Code/X86CodetreeToICode.ML index bf49f44b..2c33ad34 100644 --- a/mlsource/MLCompiler/CodeTree/X86Code/X86CodetreeToICode.ML +++ b/mlsource/MLCompiler/CodeTree/X86Code/X86CodetreeToICode.ML @@ -1,3968 +1,3975 @@ (* - Copyright David C. J. Matthews 2016-18 + Copyright David C. J. Matthews 2016-20 This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) functor X86CodetreeToICode( structure BACKENDTREE: BackendIntermediateCodeSig structure ICODE: ICodeSig structure DEBUG: DEBUGSIG structure X86FOREIGN: FOREIGNCALLSIG structure ICODETRANSFORM: X86ICODETRANSFORMSIG structure CODE_ARRAY: CODEARRAYSIG sharing ICODE.Sharing = ICODETRANSFORM.Sharing = CODE_ARRAY.Sharing ): GENCODESIG = struct open BACKENDTREE open Address open ICODE open CODE_ARRAY exception InternalError = Misc.InternalError local val regs = case targetArch of Native32Bit => [eax, ebx] | Native64Bit => [eax, ebx, r8, r9, r10] | ObjectId32Bit => [eax, esi, r8, r9, r10] val fpResult = case targetArch of Native32Bit => FPReg fp0 | _ => XMMReg xmm0 val fpArgRegs = case targetArch of Native32Bit => [] | _ => [xmm0, xmm1, xmm2] in val generalArgRegs = List.map GenReg regs val floatingPtArgRegs = List.map XMMReg fpArgRegs fun resultReg GeneralType = GenReg eax | resultReg DoubleFloatType = fpResult | resultReg SingleFloatType = fpResult end (* tag a short constant *) fun tag c = 2 * c + 1 (* shift a short constant, but don't set tag bit *) fun semitag c = 2 * c (* Reverse a list and append the second. This is used a lot when converting between the reverse and forward list versions. e.g. codeToICode and codeToICodeRev *) fun revApp([], l) = l | revApp(hd :: tl, l) = revApp(tl, hd :: l) datatype blockStruct = BlockSimple of x86ICode | BlockExit of x86ICode | BlockLabel of int | BlockFlow of controlFlow | BlockBegin of { regArgs: (preg * reg) list, stackArgs: stackLocn list } | BlockRaiseAndHandle of x86ICode * int | BlockOptionalHandle of {call: x86ICode, handler: int, label: int } local open RunCall val F_mutable_bytes = Word.fromLargeWord(Word8.toLargeWord(Word8.orb (F_mutable, F_bytes))) fun makeRealConst l = let val r = allocateByteMemory(0wx8 div bytesPerWord, F_mutable_bytes) fun setBytes([], _) = () | setBytes(hd::tl, n) = (storeByte(r, n, hd); setBytes(tl, n+0wx1)) val () = setBytes(l, 0w0) val () = clearMutableBit r in r end in (* These are floating point constants used to change and mask the sign bit. *) val realSignBit: machineWord = makeRealConst [0wx00, 0wx00, 0wx00, 0wx00, 0wx00, 0wx00, 0wx00, 0wx80] and realAbsMask: machineWord = makeRealConst [0wxff, 0wxff, 0wxff, 0wxff, 0wxff, 0wxff, 0wxff, 0wx7f] and floatSignBit: machineWord = makeRealConst [0wx00, 0wx00, 0wx00, 0wx80, 0wx00, 0wx00, 0wx00, 0wx00] and floatAbsMask: machineWord = makeRealConst [0wxff, 0wxff, 0wxff, 0wx7f, 0wx00, 0wx00, 0wx00, 0wx00] end datatype commutative = Commutative | NonCommutative (* Check that a large-word constant looks right and get the value as a large int*) fun largeWordConstant value = if isShort value then raise InternalError "largeWordConstant: invalid" else let val addr = toAddress value in if length addr <> nativeWordSize div wordSize orelse flags addr <> F_bytes then raise InternalError "largeWordConstant: invalid" else (); LargeWord.toLargeInt(RunCall.unsafeCast addr) end fun codeFunctionToX86({body, localCount, name, argTypes, resultType=fnResultType, closure, ...}:bicLambdaForm, debugSwitches, resultClosure) = let (* Pseudo-registers are allocated sequentially and the properties added to the list. *) val pregCounter = ref 0 val pregPropList = ref [] fun newPReg() = let val regNo = !pregCounter before pregCounter := !pregCounter + 1 val () = pregPropList := RegPropGeneral :: !pregPropList in PReg regNo end and newUReg() = let val regNo = !pregCounter before pregCounter := !pregCounter + 1 val () = pregPropList := RegPropUntagged :: !pregPropList in PReg regNo end and newStackLoc size = let val regNo = !pregCounter before pregCounter := !pregCounter + 1 val () = pregPropList := RegPropStack size :: !pregPropList in StackLoc{size=size, rno=regNo} end and newMergeReg() = let val regNo = !pregCounter before pregCounter := !pregCounter + 1 val () = pregPropList := RegPropMultiple :: !pregPropList in PReg regNo end datatype locationValue = NoLocation | PregLocation of preg | ContainerLocation of { container: stackLocn, stackOffset: int } val locToPregArray = Array.array(localCount, NoLocation) val labelCounter = ref 1 (* Start at 1. Zero is used for the root. *) fun newLabel() = !labelCounter before labelCounter := !labelCounter + 1 val ccRefCounter = ref 0 fun newCCRef() = CcRef(!ccRefCounter) before ccRefCounter := !ccRefCounter + 1 fun constantAsArgument value = if isShort value then IntegerConstant(tag(Word.toLargeIntX(toShort value))) else AddressConstant value (* Create the branch condition from the test, isSigned and jumpOn values. (In)equality tests are the same for signed and unsigned values. *) local open BuiltIns in fun testAsBranch(TestEqual, _, true) = JE | testAsBranch(TestEqual, _, false) = JNE (* Signed tests *) | testAsBranch(TestLess, true, true) = JL | testAsBranch(TestLess, true, false) = JGE | testAsBranch(TestLessEqual, true, true) = JLE | testAsBranch(TestLessEqual, true, false) = JG | testAsBranch(TestGreater, true, true) = JG | testAsBranch(TestGreater, true, false) = JLE | testAsBranch(TestGreaterEqual, true, true) = JGE | testAsBranch(TestGreaterEqual, true, false) = JL (* Unsigned tests *) | testAsBranch(TestLess, false, true) = JB | testAsBranch(TestLess, false, false) = JNB | testAsBranch(TestLessEqual, false, true) = JNA | testAsBranch(TestLessEqual, false, false) = JA | testAsBranch(TestGreater, false, true) = JA | testAsBranch(TestGreater, false, false) = JNA | testAsBranch(TestGreaterEqual, false, true) = JNB | testAsBranch(TestGreaterEqual, false, false) = JB | testAsBranch(TestUnordered, _, _) = raise InternalError "TestUnordered" (* Switch the direction of a test if we turn c op x into x op c. *) fun leftRightTest TestEqual = TestEqual | leftRightTest TestLess = TestGreater | leftRightTest TestLessEqual = TestGreaterEqual | leftRightTest TestGreater = TestLess | leftRightTest TestGreaterEqual = TestLessEqual | leftRightTest TestUnordered = TestUnordered end (* Overflow check. This raises Overflow if the overflow bit is set in the cc. This generates a single block for the function unless there is a handler. As well as reducing the size of the code this also means that overflow checks are generally JO instructions to the end of the code. Since the default branch prediction is not to take forward jumps this should improve prefetching on the normal, non-overflow, path. *) fun checkOverflow ({currHandler=NONE, overflowBlock=ref(SOME overFlowLab), ...}) ccRef = (* It's already been set and there's no surrounding handler - use this. *) let val noOverflowLab = newLabel() in [ BlockFlow(Conditional{ ccRef=ccRef, condition=JO, trueJump=overFlowLab, falseJump=noOverflowLab }), BlockLabel noOverflowLab ] end | checkOverflow ({currHandler=NONE, overflowBlock, ...}) ccRef = let (* *) val overFlowLab = newLabel() and noOverflowLab = newLabel() val packetReg = newPReg() val () = overflowBlock := SOME overFlowLab in [ BlockFlow(Conditional{ ccRef=ccRef, condition=JO, trueJump=overFlowLab, falseJump=noOverflowLab }), BlockLabel overFlowLab, BlockSimple(LoadArgument{source=AddressConstant(toMachineWord(Overflow)), dest=packetReg, kind=movePolyWord}), BlockExit(RaiseExceptionPacket{packetReg=packetReg}), BlockLabel noOverflowLab ] end | checkOverflow ({currHandler=SOME h, ...}) ccRef = let val overFlowLab = newLabel() and noOverflowLab = newLabel() val packetReg = newPReg() in [ BlockFlow(Conditional{ ccRef=ccRef, condition=JO, trueJump=overFlowLab, falseJump=noOverflowLab }), BlockLabel overFlowLab, BlockSimple(LoadArgument{source=AddressConstant(toMachineWord(Overflow)), dest=packetReg, kind=movePolyWord}), BlockRaiseAndHandle(RaiseExceptionPacket{packetReg=packetReg}, h), BlockLabel noOverflowLab ] end fun setAndRestoreRounding (rndMode, doWithRounding) = let open IEEEReal val savedRnd = newUReg() and setRnd = newUReg() in case fpMode of FPModeX87 => [BlockSimple(GetX87ControlReg{dest=savedRnd})] @ (* Set the appropriate bits in the control word. *) (case rndMode of TO_NEAREST => (* The bits need to be zero - just mask them. *) [BlockSimple( ArithmeticFunction{oper=AND, resultReg=setRnd, operand1=savedRnd, operand2=IntegerConstant 0xf3ff, ccRef=newCCRef(), opSize=OpSize32})] | TO_NEGINF => let val wrk = newUReg() in (* Mask the bits and set to 01 *) [BlockSimple( ArithmeticFunction{oper=AND, resultReg=wrk, operand1=savedRnd, operand2=IntegerConstant 0xf3ff, ccRef=newCCRef(), opSize=OpSize32}), BlockSimple( ArithmeticFunction{oper=OR, resultReg=setRnd, operand1=savedRnd, operand2=IntegerConstant 0x400, ccRef=newCCRef(), opSize=OpSize32})] end | TO_POSINF => let val wrk = newUReg() in (* Mask the bits and set to 10 *) [BlockSimple( ArithmeticFunction{oper=AND, resultReg=wrk, operand1=savedRnd, operand2=IntegerConstant 0xf3ff, ccRef=newCCRef(), opSize=OpSize32}), BlockSimple( ArithmeticFunction{oper=OR, resultReg=setRnd, operand1=savedRnd, operand2=IntegerConstant 0x800, ccRef=newCCRef(), opSize=OpSize32})] end | TO_ZERO => (* The bits need to be one - just set them. *) [BlockSimple( ArithmeticFunction{oper=OR, resultReg=setRnd, operand1=savedRnd, operand2=IntegerConstant 0xc00, ccRef=newCCRef(), opSize=OpSize32})]) @ [BlockSimple(SetX87ControlReg{source=setRnd})] @ doWithRounding() @ (* Restore the original rounding. *) [BlockSimple(SetX87ControlReg{source=savedRnd})] | FPModeSSE2 => [BlockSimple(GetSSE2ControlReg{dest=savedRnd})] @ (* Set the appropriate bits in the control word. *) (case rndMode of TO_NEAREST => (* The bits need to be zero - just mask them. *) [BlockSimple( ArithmeticFunction{oper=AND, resultReg=setRnd, operand1=savedRnd, operand2=IntegerConstant 0xffff9fff, ccRef=newCCRef(), opSize=OpSize32})] | TO_NEGINF => let val wrk = newUReg() in (* Mask the bits and set to 01 *) [BlockSimple( ArithmeticFunction{oper=AND, resultReg=wrk, operand1=savedRnd, operand2=IntegerConstant 0xffff9fff, ccRef=newCCRef(), opSize=OpSize32}), BlockSimple( ArithmeticFunction{oper=OR, resultReg=setRnd, operand1=savedRnd, operand2=IntegerConstant 0x2000, ccRef=newCCRef(), opSize=OpSize32})] end | TO_POSINF => let val wrk = newUReg() in (* Mask the bits and set to 10 *) [BlockSimple( ArithmeticFunction{oper=AND, resultReg=wrk, operand1=savedRnd, operand2=IntegerConstant 0xffff9fff, ccRef=newCCRef(), opSize=OpSize32}), BlockSimple( ArithmeticFunction{oper=OR, resultReg=setRnd, operand1=savedRnd, operand2=IntegerConstant 0x4000, ccRef=newCCRef(), opSize=OpSize32})] end | TO_ZERO => (* The bits need to be one - just set them. *) [BlockSimple( ArithmeticFunction{oper=OR, resultReg=setRnd, operand1=savedRnd, operand2=IntegerConstant 0x6000, ccRef=newCCRef(), opSize=OpSize32})]) @ [BlockSimple(SetSSE2ControlReg{source=setRnd})] @ doWithRounding() @ [BlockSimple(SetSSE2ControlReg{source=savedRnd})] end (* Put a floating point value into a box or tag it so the value can be held in a general register. *) fun boxOrTagReal(srcReg, destReg, precision) = if precision = BuiltIns.PrecDouble orelse wordSize <> 0w8 then let open BuiltIns val boxFloat = case (fpMode, precision) of (FPModeX87, PrecDouble) => BoxX87Double | (FPModeX87, PrecSingle) => BoxX87Float | (FPModeSSE2, PrecDouble) => BoxSSE2Double | (FPModeSSE2, PrecSingle) => BoxSSE2Float in [BlockSimple(BoxValue{boxKind=boxFloat, source=srcReg, dest=destReg, saveRegs=[]})] end else [BlockSimple(TagFloat{source=srcReg, dest=destReg})] (* Indicate that the base address is actually an object index where appropriate. *) val memIndexOrObject = case targetArch of ObjectId32Bit => ObjectIndex | _ => NoMemIndex (* Generally we have an offset in words and no index register. *) fun wordOffsetAddress(offset, baseReg: preg): argument = MemoryLocation{offset=offset*Word.toInt wordSize, base=baseReg, index=memIndexOrObject, cache=NONE} (* The large-word operations all work on the value within the box pointed at by the register. We generate all large-word operations using this even where the X86 instruction requires a register. This allows the next level to optimise cases of cascaded instructions and avoid creating boxes for intermediate values. *) fun wordAt reg = wordOffsetAddress(0, reg) val returnAddressEntry = newStackLoc 1 datatype argLoc = ArgInReg of { realReg: reg, argReg: preg } | ArgOnStack of { stackOffset: int, stackReg: stackLocn } (* Pseudo-regs for the result, the closure and the args that were passed in real regs. *) val resultTarget = newPReg() val closureRegAddr = newPReg() (* Create a map for the arguments indicating their register or stack location. *) local (* Select the appropriate argument register depending on the argument type. *) fun argTypesToArgEntries([], _, _, _) = ([], [], [], []) | argTypesToArgEntries(DoubleFloatType :: tl, gRegs, fpReg :: fpRegs, n) = let val (argTypes, argCode, argRegs, stackArgs) = argTypesToArgEntries(tl, gRegs, fpRegs, n-1) val pRegArg = newPReg() and uRegArg = newUReg() in (ArgInReg{realReg=fpReg, argReg=pRegArg} :: argTypes, boxOrTagReal(uRegArg, pRegArg, BuiltIns.PrecDouble) @ argCode, (uRegArg, fpReg) :: argRegs, stackArgs) end | argTypesToArgEntries(SingleFloatType :: tl, gRegs, fpReg :: fpRegs, n) = let val (argTypes, argCode, argRegs, stackArgs) = argTypesToArgEntries(tl, gRegs, fpRegs, n-1) val pRegArg = newPReg() and uRegArg = newUReg() in (ArgInReg{realReg=fpReg, argReg=pRegArg} :: argTypes, boxOrTagReal(uRegArg, pRegArg, BuiltIns.PrecSingle) @ argCode, (uRegArg, fpReg) :: argRegs, stackArgs) end | argTypesToArgEntries(_ :: tl, gReg :: gRegs, fpRegs, n) = (* This deals with general arguments but also with extra floating point arguments. They are boxed as usual. *) let val (argTypes, argCode, argRegs, stackArgs) = argTypesToArgEntries(tl, gRegs, fpRegs, n-1) val argReg=newPReg() in (ArgInReg{realReg=gReg, argReg=argReg} :: argTypes, argCode, (argReg, gReg) :: argRegs, stackArgs) end | argTypesToArgEntries(_ :: tl, [], fpRegs, n) = let val (argTypes, argCode, argRegs, stackArgs) = argTypesToArgEntries(tl, [], fpRegs, n-1) val stackLoc = newStackLoc 1 in (ArgOnStack {stackOffset=n, stackReg = stackLoc } :: argTypes, argCode, argRegs, stackLoc :: stackArgs) end val (argEntries, argCode, argRegs, stackArguments) = argTypesToArgEntries(argTypes, generalArgRegs, floatingPtArgRegs, List.length argTypes) val clReg = case closure of [] => [] | _ => [(closureRegAddr, GenReg edx)] in val argumentVector = Vector.fromList argEntries (* Start code for the function. *) val beginInstructions = argCode @ [BlockBegin{regArgs=clReg @ argRegs, stackArgs=stackArguments @ [returnAddressEntry]}] (* The number of arguments on the stack. Needed in return instrs and tail calls. *) val currentStackArgs = List.length stackArguments end (* The return instruction. This can be added on to various tails but there is always one at the end anyway. *) fun returnInstruction({stackPtr, ...}, target, tailCode) = let val (returnCode, resReg) = case fnResultType of GeneralType => ([], target) | DoubleFloatType => let val resReg = newUReg() in ([BlockSimple(LoadArgument{source=wordAt target, dest=resReg, kind=MoveDouble})], resReg) end | SingleFloatType => let val resReg = newUReg() val unpack = if wordSize = 0w8 then BlockSimple(UntagFloat{source=RegisterArgument target, dest=resReg, cache=NONE}) else BlockSimple(LoadArgument{source=wordAt target, dest=resReg, kind=MoveFloat}) in ([unpack], resReg) end in BlockExit(ReturnResultFromFunction{resultReg=resReg, realReg=resultReg fnResultType, numStackArgs=currentStackArgs}) :: returnCode @ (if stackPtr <> 0 then BlockSimple(ResetStackPtr{numWords=stackPtr, preserveCC=false}) :: tailCode else tailCode) end (* This controls what codeAsArgument returns. Different instructions have different requirements. If an option is set to false the value is instead loaded into a new preg. "const32s" means that it will fit into 32-bits. Any constant satisfies that on X86/32 but on the X86/64 we don't allow addresses because we can't be sure whether they will fit or not. *) type allowedArgument = { anyConstant: bool, const32s: bool, memAddr: bool, existingPreg: bool } val allowInMemMove = (* We can move a 32-bit constant into memory but not a long constant. *) { anyConstant=false, const32s=true, memAddr=false, existingPreg=true } and allowInPReg = { anyConstant=false, const32s=false, memAddr=false, existingPreg=true } (* AllowDefer can be used to ensure that any side-effects are done before something else but otherwise we only evaluate afterwards. *) and allowDefer = { anyConstant=true, const32s=true, memAddr=true, existingPreg=true } datatype destination = SpecificPReg of preg | NoResult | Allowed of allowedArgument (* Context type. *) type context = { loopArgs: (preg list * int * int) option, stackPtr: int, currHandler: int option, overflowBlock: int option ref } (* If a preg has been provided, use that, otherwise generate a new one. *) fun asTarget(SpecificPReg preg) = preg | asTarget NoResult = newPReg() | asTarget(Allowed _) = newPReg() fun moveIfNotAllowed(NoResult, code, arg) = (code, arg, false) | moveIfNotAllowed(Allowed{anyConstant=true, ...}, code, arg as AddressConstant _) = (code, arg, false) | moveIfNotAllowed(Allowed{anyConstant=true, ...}, code, arg as IntegerConstant _) = (code, arg, false) | moveIfNotAllowed(dest as Allowed{const32s=true, ...}, code, arg as IntegerConstant value) = (* This is allowed if the value is within 32-bits *) if is32bit value then (code, arg, false) else moveToTarget(dest, code, arg) | moveIfNotAllowed(dest as Allowed{const32s=true, ...}, code, arg as AddressConstant _) = if targetArch = Native32Bit then (code, arg, false) (* We can store the address directly *) else moveToTarget(dest, code, arg) | moveIfNotAllowed(Allowed{existingPreg=true, ...}, code, arg as RegisterArgument(PReg _)) = (code, arg, false) | moveIfNotAllowed(Allowed{memAddr=true, ...}, code, arg as MemoryLocation _) = (code, arg, false) | moveIfNotAllowed(dest, code, arg) = moveToTarget(dest, code, arg) and moveToTarget(dest, code, arg) = let val target = asTarget dest val moveSize = case arg of AddressConstant _ => movePolyWord | MemoryLocation _ => movePolyWord | _ => moveNativeWord in (code @ [BlockSimple(LoadArgument{source=arg, dest=target, kind=moveSize})], RegisterArgument target, false) end (* Create a bool result from a test by returning true or false. *) fun makeBoolResultRev(condition, ccRef, target, testCode) = let val trueLab = newLabel() and falseLab = newLabel() and mergeLab = newLabel() val mergeReg = newMergeReg() in BlockSimple(LoadArgument{dest=target, source=RegisterArgument mergeReg, kind=Move32Bit}) :: BlockLabel mergeLab :: BlockFlow(Unconditional mergeLab) :: BlockSimple(LoadArgument{dest=mergeReg, source=IntegerConstant(tag 0), kind=Move32Bit}) :: BlockLabel falseLab :: BlockFlow(Unconditional mergeLab) :: BlockSimple(LoadArgument{dest=mergeReg, source=IntegerConstant(tag 1), kind=Move32Bit}) :: BlockLabel trueLab :: BlockFlow(Conditional{ ccRef=ccRef, condition=condition, trueJump=trueLab, falseJump=falseLab }) :: testCode end fun moveIfNotAllowedRev(NoResult, code, arg) = (code, arg, false) | moveIfNotAllowedRev(Allowed{anyConstant=true, ...}, code, arg as AddressConstant _) = (code, arg, false) | moveIfNotAllowedRev(Allowed{anyConstant=true, ...}, code, arg as IntegerConstant _) = (code, arg, false) | moveIfNotAllowedRev(dest as Allowed{const32s=true, ...}, code, arg as IntegerConstant value) = (* This is allowed if the value is within 32-bits *) if is32bit value then (code, arg, false) else moveToTargetRev(dest, code, arg) | moveIfNotAllowedRev(dest as Allowed{const32s=true, ...}, code, arg as AddressConstant _) = if targetArch = Native32Bit then (code, arg, false) else moveToTargetRev(dest, code, arg) | moveIfNotAllowedRev(Allowed{existingPreg=true, ...}, code, arg as RegisterArgument(PReg _)) = (code, arg, false) | moveIfNotAllowedRev(Allowed{memAddr=true, ...}, code, arg as MemoryLocation _) = (code, arg, false) | moveIfNotAllowedRev(dest, code, arg) = moveToTargetRev(dest, code, arg) and moveToTargetRev(dest, code, arg) = let val target = asTarget dest val moveSize = case arg of AddressConstant _ => movePolyWord | MemoryLocation _ => movePolyWord | _ => moveNativeWord in (BlockSimple(LoadArgument{source=arg, dest=target, kind=moveSize}) :: code, RegisterArgument target, false) end (* Use a move if there's no offset or index. We could use an add if there's no index. *) and loadAddress{base, offset=0, index=NoMemIndex, dest} = LoadArgument{source=RegisterArgument base, dest=dest, kind=movePolyWord} | loadAddress{base, offset, index, dest} = LoadEffectiveAddress{base=SOME base, offset=offset, dest=dest, index=index, opSize=nativeWordOpSize} and codeToICodeTarget(instr, context: context, isTail, target) = (* This is really for backwards compatibility. *) let val (code, _, _) = codeToICode(instr, context, isTail, SpecificPReg target) in code end and codeToPReg(instr, context) = let (* Many instructions require an argument in a register. If it's already in a register use that rather than creating a new one. *) val (code, result, _) = codeToICode(instr, context, false, Allowed allowInPReg) val preg = case result of RegisterArgument pr => pr | _ => raise InternalError "codeToPReg" in (code, preg) end and codeToPRegRev(instr, context, tailCode) = let (* Many instructions require an argument in a register. If it's already in a register use that rather than creating a new one. *) val (code, result, _) = codeToICodeRev(instr, context, false, Allowed allowInPReg, tailCode) val preg = case result of RegisterArgument pr => pr | _ => raise InternalError "codeToPRegRev" in (code, preg) end and codeToICode(instr, context, isTail, destination) = let val (code, dest, haveExited) = codeToICodeRev(instr, context, isTail, destination, []) in (List.rev code, dest, haveExited) end (* Main function to turn the codetree into ICode. Optimisation is generally left to later passes. This does detect tail recursion. This builds the result up in reverse order. There was an allocation hotspot in loadFields in the BICTuple case which was eliminated by building the list in reverse and then reversing the result. It seems better to build the list in reverse generally but for the moment there are too many special cases to do everything. *) and codeToICodeRev(BICNewenv (bindings, exp), context: context as {stackPtr=initialSp, ...} , isTail, destination, tailCode) = let (* Process a list of bindings. We need to accumulate the space used by any containers and reset the stack pointer at the end if necessary. *) fun doBindings([], context, tailCode) = (tailCode, context) | doBindings(BICDeclar{value=BICExtract(BICLoadLocal l), addr, ...} :: decs, context, tailCode) = let (* Giving a new name to an existing entry. This should have been removed at a higher level but it doesn't always seem to be. In particular we must treat this specially if it's a container. *) val original = Array.sub(locToPregArray, l) val () = Array.update(locToPregArray, addr, original) in doBindings(decs, context, tailCode) end | doBindings(BICDeclar{value, addr, ...} :: decs, context, tailCode) = let val (code, dest) = codeToPRegRev(value, context, tailCode) val () = Array.update(locToPregArray, addr, PregLocation dest) in doBindings(decs, context, code) end | doBindings(BICRecDecs [{lambda, addr, ...}] :: decs, context, tailCode) = (* We shouldn't have single entries in RecDecs but it seems to occur at the moment. *) let val dest = newPReg() val (code, _, _) = codeToICodeRev(BICLambda lambda, context, false, SpecificPReg dest, tailCode) val () = Array.update(locToPregArray, addr, PregLocation dest) in doBindings(decs, context, code) end | doBindings(BICRecDecs recDecs :: decs, context, tailCode) = let val destRegs = map (fn _ => newPReg()) recDecs (* First build the closures as mutable cells containing zeros. Set the entry in the address table to the register containing the address. *) fun makeClosure({lambda={closure, ...}, addr, ...}, dest, c) = let val () = Array.update(locToPregArray, addr, PregLocation dest) val sizeClosure = List.length closure + (if targetArch = ObjectId32Bit then 2 else 1) open Address fun clear n = if n = sizeClosure then [BlockSimple(AllocateMemoryOperation{size=sizeClosure, flags=if targetArch = ObjectId32Bit then Word8.orb(F_mutable, F_closure) else F_mutable, dest=dest, saveRegs=[]})] else (clear (n+1) @ [BlockSimple( StoreArgument{source=IntegerConstant(tag 0), base=dest, offset=n*Word.toInt wordSize, index=memIndexOrObject, kind=movePolyWord, isMutable=false})]) in c @ clear 0 @ [BlockSimple InitialisationComplete] end val allocClosures = ListPair.foldlEq makeClosure [] (recDecs, destRegs) fun setClosure({lambda as {closure, ...}, ...}, dest, l) = let val clResult = makeConstantClosure() val () = codeFunctionToX86(lambda, debugSwitches, clResult) (* Basically the same as tuple except we load the address of the closure we've made. *) fun loadFields([], _) = [] | loadFields(f :: rest, n) = let val (code, source, _) = codeToICode(BICExtract f, context, false, Allowed allowInMemMove) val storeValue = [BlockSimple(StoreArgument{ source=source, base=dest, offset=n*Word.toInt wordSize, index=memIndexOrObject, kind=movePolyWord, isMutable=false })] in code @ storeValue @ loadFields(rest, n+1) end val setCodeAddress = if targetArch = ObjectId32Bit then let (* We can't get the code address until run time. *) val codeReg = newUReg() val closureReg = newPReg() in map BlockSimple [ LoadArgument{ source=AddressConstant(toMachineWord clResult), dest=closureReg, kind=movePolyWord}, LoadArgument{ source=MemoryLocation{offset=0, base=closureReg, index=ObjectIndex, cache=NONE}, dest=codeReg, kind=Move64Bit}, StoreArgument{ source=RegisterArgument codeReg, offset=0, base=dest, index=ObjectIndex, kind=moveNativeWord, isMutable=false} ] end else let val codeAddr = codeAddressFromClosure clResult val (code, source, _) = moveIfNotAllowed(Allowed allowInMemMove, [], AddressConstant codeAddr) in code @ [BlockSimple( StoreArgument{ source=source, base=dest, offset=0, index=NoMemIndex, kind=movePolyWord, isMutable=false })] end val setFields = setCodeAddress @ loadFields(closure, if targetArch = ObjectId32Bit then 2 else 1) in l @ setFields @ [BlockSimple(LockMutable{addr=dest})] end val setClosures = ListPair.foldlEq setClosure [] (recDecs, destRegs) val code = List.rev(allocClosures @ setClosures) in doBindings(decs, context, code @ tailCode) end | doBindings(BICNullBinding exp :: decs, context, tailCode) = let val (code, _, _) = codeToICodeRev(exp, context, false, NoResult, tailCode) (* And discard result. *) in doBindings(decs, context, code) end | doBindings(BICDecContainer{ addr, size } :: decs, {loopArgs, stackPtr, currHandler, overflowBlock}, tailCode) = let val containerReg = newStackLoc size val () = Array.update(locToPregArray, addr, ContainerLocation{container=containerReg, stackOffset=stackPtr+size}) in doBindings(decs, {loopArgs=loopArgs, stackPtr=stackPtr+size, currHandler=currHandler, overflowBlock=overflowBlock}, BlockSimple(ReserveContainer{size=size, container=containerReg}) :: tailCode) end val (codeBindings, resContext as {stackPtr=finalSp, ...}) = doBindings(bindings, context, tailCode) (* If we have had a container we'll need to reset the stack *) in if initialSp <> finalSp then let val _ = finalSp >= initialSp orelse raise InternalError "codeToICode - stack ptr" val bodyReg = newPReg() and resultReg = asTarget destination val (codeExp, result, haveExited) = codeToICodeRev(exp, resContext, isTail, SpecificPReg bodyReg, codeBindings) val afterAdjustSp = if haveExited then codeExp else BlockSimple(LoadArgument{source=result, dest=resultReg, kind=movePolyWord}) :: BlockSimple(ResetStackPtr{numWords=finalSp-initialSp, preserveCC=false}) :: codeExp in (afterAdjustSp, RegisterArgument resultReg, haveExited) end else codeToICodeRev(exp, resContext, isTail, destination, codeBindings) end | codeToICodeRev(BICConstnt(value, _), _, _, destination, tailCode) = moveIfNotAllowedRev(destination, tailCode, constantAsArgument value) | codeToICodeRev(BICExtract(BICLoadLocal l), {stackPtr, ...}, _, destination, tailCode) = ( case Array.sub(locToPregArray, l) of NoLocation => raise InternalError "codeToICodeRev - local unset" | PregLocation preg => moveIfNotAllowedRev(destination, tailCode, RegisterArgument preg) | ContainerLocation{container, stackOffset} => (* This always returns a ContainerAddr whatever the "allowed". *) (tailCode, ContainerAddr{container=container, stackOffset=stackPtr-stackOffset}, false) ) | codeToICodeRev(BICExtract(BICLoadArgument a), {stackPtr, ...}, _, destination, tailCode) = ( case Vector.sub(argumentVector, a) of ArgInReg{argReg, ...} => (* It was originally in a register. It's now in a preg. *) moveIfNotAllowedRev(destination, tailCode, RegisterArgument argReg) | ArgOnStack{stackOffset, stackReg} => (* Pushed before call. *) let val target = asTarget destination in (BlockSimple(LoadArgument{ source=StackLocation{wordOffset=stackOffset+stackPtr, container=stackReg, field=0, cache=NONE}, dest=target, kind=moveNativeWord}) :: tailCode, RegisterArgument target, false) end ) | codeToICodeRev(BICExtract(BICLoadClosure c), _, _, destination, tailCode) = let (* Add the number of words for the code address. This is 1 in native but 2 in 32-in-64. *) val offset = case targetArch of ObjectId32Bit => c+2 | _ => c+1 in if c >= List.length closure then raise InternalError "BICExtract: closure" else (); (* N.B. We need to add one to the closure entry because zero is the code address. *) moveIfNotAllowedRev(destination, tailCode, wordOffsetAddress(offset, closureRegAddr)) end | codeToICodeRev(BICExtract BICLoadRecursive, _, _, destination, tailCode) = (* If the closure is empty we must use the constant. We can't guarantee that the caller will actually load the closure register if it knows the closure is empty. *) moveIfNotAllowedRev(destination, tailCode, case closure of [] => AddressConstant(closureAsAddress resultClosure) | _ => RegisterArgument closureRegAddr) | codeToICodeRev(BICField{base, offset}, context, _, destination, tailCode) = let val (codeBase, baseEntry, _) = codeToICodeRev(base, context, false, Allowed allowInPReg, tailCode) in (* This should not be used with a container. *) case baseEntry of RegisterArgument baseR => moveIfNotAllowedRev(destination, codeBase, wordOffsetAddress(offset, baseR)) | _ => raise InternalError "codeToICodeRev-BICField" end | codeToICodeRev(BICLoadContainer{base, offset}, context, _, destination, tailCode) = let val (codeBase, baseEntry, _) = codeToICodeRev(base, context, false, Allowed allowInPReg, tailCode) val multiplier = Word.toInt(nativeWordSize div wordSize) in (* If this is a local container we extract the field. *) case baseEntry of RegisterArgument baseR => moveIfNotAllowedRev(destination, codeBase, wordOffsetAddress(offset*multiplier, baseR)) | ContainerAddr{container, stackOffset} => let val target = asTarget destination val finalOffset = stackOffset+offset val _ = finalOffset >= 0 orelse raise InternalError "offset" in (BlockSimple(LoadArgument{ source=StackLocation{wordOffset=finalOffset, container=container, field=offset, cache=NONE}, dest=target, kind=moveNativeWord}) :: tailCode, RegisterArgument target, false) end | _ => raise InternalError "codeToICodeRev-BICField" end | codeToICodeRev(BICEval{function, argList, resultType, ...}, context as { currHandler, ...}, isTail, destination, tailCode) = let val target = asTarget destination (* Create pregs for the closure and each argument. *) val clPReg = newPReg() (* If we have a constant closure we can go directly to the entry point. If the closure is a single word we don't need to load the closure register. *) val (functionCode, closureEntry, callKind) = case function of BICConstnt(addr, _) => let val addrAsAddr = toAddress addr (* If this is a closure we're still compiling we can't get the code address. However if this is directly recursive we can use the recursive convention. *) in if wordEq(closureAsAddress resultClosure, addr) then (tailCode, [], Recursive) else if flags addrAsAddr <> Address.F_words andalso flags addrAsAddr <> Address.F_closure then (BlockSimple(LoadArgument{source=AddressConstant addr, dest=clPReg, kind=movePolyWord}) :: tailCode, [(RegisterArgument clPReg, GenReg edx)], FullCall) else if targetArch = ObjectId32Bit then (* We can't actually load the code address here. *) let val addrLength = length addrAsAddr val _ = addrLength >= 0w1 orelse raise InternalError "BICEval address" val _ = flags addrAsAddr = Address.F_closure orelse raise InternalError "BICEval address not a closure" in if addrLength = 0w2 then (tailCode, [], ConstantCode addr) else (BlockSimple(LoadArgument{source=AddressConstant addr, dest=clPReg, kind=movePolyWord}) :: tailCode, [(RegisterArgument clPReg, GenReg edx)], ConstantCode addr) end else (* Native 32 or 64-bits. *) let val addrLength = length addrAsAddr val _ = addrLength >= 0w1 orelse raise InternalError "BICEval address" val codeAddr = loadWord(addrAsAddr, 0w0) val _ = isCode (toAddress codeAddr) orelse raise InternalError "BICEval address not code" in if addrLength = 0w1 then (tailCode, [], ConstantCode codeAddr) else (BlockSimple(LoadArgument{source=AddressConstant addr, dest=clPReg, kind=movePolyWord}) :: tailCode, [(RegisterArgument clPReg, GenReg edx)], ConstantCode codeAddr) end end | BICExtract BICLoadRecursive => ( (* If the closure is empty we don't need to load rdx *) case closure of [] => (tailCode, [], Recursive) | _ => (BlockSimple(LoadArgument {source=RegisterArgument closureRegAddr, dest=clPReg, kind=movePolyWord}) :: tailCode, [(RegisterArgument clPReg, GenReg edx)], Recursive) ) | function => (* General case. *) (#1 (codeToICodeRev(function, context, false, SpecificPReg clPReg, tailCode)), [(RegisterArgument clPReg, GenReg edx)], FullCall) (* Optimise arguments. We have to be careful with tail-recursive functions because they need to save any stack arguments that could be overwritten. This is complicated because we overwrite the stack before loading the register arguments. In some circumstances it could be safe but for the moment leave it. This should be safe in the new code-transform but not the old codeICode. Currently we don't allow memory arguments at all. There's the potential for problems later. Memory arguments could possibly lead to aliasing of the stack if the memory actually refers to a container on the stack. That would mess up the code that ensures that stack arguments are stored in the right order. *) (* We don't allow long constants in stack arguments to a tail-recursive call because we may use a memory move to set them. We also don't allow them in 32-in-64 because we can't push an address constant. *) val allowInStackArg = Allowed {anyConstant=not isTail andalso targetArch <> ObjectId32Bit, const32s=true, memAddr=false, existingPreg=not isTail } and allowInRegArg = Allowed {anyConstant=true, const32s=true, memAddr=false, existingPreg=not isTail } (* Load the first arguments into registers and the rest to the stack. *) fun loadArgs ([], _, _, tailCode) = (tailCode, [], []) | loadArgs ((arg, DoubleFloatType) :: args, gRegs, fpReg :: fpRegs, tailCode) = let (* Floating point register argument. *) val (c, r) = codeToPRegRev(arg, context, tailCode) val r1 = newUReg() val c1 = BlockSimple(LoadArgument{source=wordAt r, dest=r1, kind=MoveDouble}) :: c val (code, regArgs, stackArgs) = loadArgs(args, gRegs, fpRegs, c1) in (code, (RegisterArgument r1, fpReg) :: regArgs, stackArgs) end | loadArgs ((arg, SingleFloatType) :: args, gRegs, fpReg :: fpRegs, tailCode) = let (* Floating point register argument. *) val (c, r) = codeToPRegRev(arg, context, tailCode) val r1 = newUReg() val c1 = if wordSize = 0w8 then BlockSimple(UntagFloat{source=RegisterArgument r, dest=r1, cache=NONE}) :: c else BlockSimple(LoadArgument{source=wordAt r, dest=r1, kind=MoveFloat}) :: c val (code, regArgs, stackArgs) = loadArgs(args, gRegs, fpRegs, c1) in (code, (RegisterArgument r1, fpReg) :: regArgs, stackArgs) end | loadArgs ((arg, _) :: args, gReg::gRegs, fpRegs, tailCode) = let (* General register argument. *) val (c, r, _) = codeToICodeRev(arg, context, false, allowInRegArg, tailCode) val (code, regArgs, stackArgs) = loadArgs(args, gRegs, fpRegs, c) in (code, (r, gReg) :: regArgs, stackArgs) end | loadArgs ((arg, _) :: args, [], fpRegs, tailCode) = let (* Stack argument. *) val (c, r, _) = codeToICodeRev(arg, context, false, allowInStackArg, tailCode) val (code, regArgs, stackArgs) = loadArgs(args, [], fpRegs, c) in (code, regArgs, r :: stackArgs) end val (codeArgs, regArgs, stackArgs) = loadArgs(argList, generalArgRegs, floatingPtArgRegs, functionCode) (* If this is at the end of the function and the result types are the same we can use a tail-recursive call. *) val tailCall = isTail andalso resultType = fnResultType val callCode = if tailCall then let val {stackPtr, ...} = context (* The number of arguments currently on the stack. *) val currentStackArgCount = currentStackArgs val newStackArgCount = List.length stackArgs (* The offset of the first argument or the return address if there are no stack arguments. N.B. We actually have currentStackArgCount+1 items on the stack including the return address. Offsets can be negative. *) val stackOffset = stackPtr val firstArgumentAddr = currentStackArgCount fun makeStackArgs([], _) = [] | makeStackArgs(arg::args, offset) = {src=arg, stack=offset} :: makeStackArgs(args, offset-1) val stackArgs = makeStackArgs(stackArgs, firstArgumentAddr) (* The stack adjustment needed to compensate for any items that have been pushed and the differences in the number of arguments. May be positive or negative. This is also the destination address of the return address so when we enter the new function the return address will be the first item on the stack. *) val stackAdjust = firstArgumentAddr - newStackArgCount (* Add an entry for the return address to the stack arguments. *) val returnEntry = {src=StackLocation{wordOffset=stackPtr, container=returnAddressEntry, field=0, cache=NONE}, stack=stackAdjust} (* Because we're storing into the stack we may be overwriting values we want. If the source of any value is a stack location below the current stack pointer we load it except in the special case where the destination is the same as the source (which is often the case with the return address). *) local fun loadArgs [] = ([], []) | loadArgs (arg :: rest) = let val (loadCode, loadedArgs) = loadArgs rest in case arg of {src as StackLocation{wordOffset, ...}, stack} => if wordOffset = stack+stackOffset (* Same location *) orelse stack+stackOffset < 0 (* Storing above current top of stack *) orelse stackOffset+wordOffset > ~ stackAdjust (* Above the last argument *) then (loadCode, arg :: loadedArgs) else let val preg = newPReg() in (BlockSimple(LoadArgument{source=src, dest=preg, kind=moveNativeWord}) :: loadCode, {src=RegisterArgument preg, stack=stack} :: loadedArgs) end | _ => (loadCode, arg :: loadedArgs) end in val (loadStackArgs, loadedStackArgs) = loadArgs(returnEntry :: stackArgs) end in BlockExit(TailRecursiveCall{regArgs=closureEntry @ regArgs, stackArgs=loadedStackArgs, stackAdjust = stackAdjust, currStackSize=stackOffset, callKind=callKind, workReg=newPReg()}) :: loadStackArgs @ codeArgs end else let val (moveResult, resReg) = case resultType of GeneralType => ([], target) | DoubleFloatType => let val fpRegDest = newUReg() in (boxOrTagReal(fpRegDest, target, BuiltIns.PrecDouble), fpRegDest) end | SingleFloatType => let val fpRegDest = newUReg() in (boxOrTagReal(fpRegDest, target, BuiltIns.PrecSingle), fpRegDest) end val call = FunctionCall{regArgs=closureEntry @ regArgs, stackArgs=stackArgs, dest=resReg, realDest=resultReg resultType, callKind=callKind, saveRegs=[]} val callBlock = case currHandler of NONE => BlockSimple call :: codeArgs | SOME h => BlockOptionalHandle{call=call, handler=h, label=newLabel()} :: codeArgs in moveResult @ callBlock end in (callCode, RegisterArgument target, tailCall (* We've exited if this was a tail jump *)) end | codeToICodeRev(BICGetThreadId, _, _, destination, tailCode) = (* Get the ID of the current thread. *) let val target = asTarget destination in (BlockSimple(LoadMemReg{offset=memRegThreadSelf, dest=target}) :: tailCode, RegisterArgument target, false) end | codeToICodeRev(BICUnary instr, context, isTail, destination, tailCode) = codeToICodeUnaryRev(instr, context, isTail, destination, tailCode) | codeToICodeRev(BICBinary instr, context, isTail, destination, tailCode) = codeToICodeBinaryRev(instr, context, isTail, destination, tailCode) | codeToICodeRev(BICArbitrary{oper, shortCond, arg1, arg2, longCall}, context, _, destination, tailCode) = let val startLong = newLabel() and resultLabel = newLabel() val target = asTarget destination val condResult = newMergeReg() (* Overflow check - if there's an overflow jump to the long precision case. *) fun jumpOnOverflow ccRef = let val noOverFlow = newLabel() in [BlockFlow(Conditional{ ccRef=ccRef, condition=JO, trueJump=startLong, falseJump=noOverFlow }), BlockLabel noOverFlow] end val (longCode, _, _) = codeToICode(longCall, context, false, SpecificPReg condResult) (* We could use a tail jump here if this is a tail. *) val (code, dest, haveExited) = ( (* Test the tag bits and skip to the long case if either is clear. *) List.rev(codeConditionRev(shortCond, context, false, startLong, [])) @ (* Try evaluating as fixed precision and jump if we get an overflow. *) codeFixedPrecisionArith(oper, arg1, arg2, context, condResult, jumpOnOverflow) @ (* If we haven't had an overflow jump to the result. *) [BlockFlow(Unconditional resultLabel), (* If we need to use the full long-precision call we come here. *) BlockLabel startLong] @ longCode @ [BlockLabel resultLabel, BlockSimple(LoadArgument{source=RegisterArgument condResult, dest=target, kind=movePolyWord})], RegisterArgument target, false) in (revApp(code, tailCode), dest, haveExited) end | codeToICodeRev(BICAllocateWordMemory instr, context, isTail, destination, tailCode) = let val (code, dest, haveExited) = codeToICodeAllocate(instr, context, isTail, destination) in (revApp(code, tailCode), dest, haveExited) end | codeToICodeRev(BICLambda(lambda as { closure = [], ...}), _, _, destination, tailCode) = (* Empty closure - create a constant closure for any recursive calls. *) let val closure = makeConstantClosure() val () = codeFunctionToX86(lambda, debugSwitches, closure) (* Return the closure itself as the value. *) in moveIfNotAllowedRev(destination, tailCode, AddressConstant(closureAsAddress closure)) end | codeToICodeRev(BICLambda(lambda as { closure, ...}), context, isTail, destination, tailCode) = (* Non-empty closure. Ignore stack closure option at the moment. *) let val closureRef = makeConstantClosure() val () = codeFunctionToX86(lambda, debugSwitches, closureRef) in if targetArch = ObjectId32Bit then let val target = asTarget destination val memAddr = newPReg() fun loadFields([], n, tlCode) = let val codeReg = newUReg() val closureReg = newPReg() in (* The code address occupies the first native word but we need to extract it at run-time. We don't currently have a way to have 64-bit constants. *) BlockSimple( StoreArgument{ source=RegisterArgument codeReg, offset=0, base=memAddr, index=ObjectIndex, kind=moveNativeWord, isMutable=false}) :: BlockSimple(LoadArgument{ source=MemoryLocation{offset=0, base=closureReg, index=ObjectIndex, cache=NONE}, dest=codeReg, kind=Move64Bit}) :: BlockSimple(LoadArgument{ source=AddressConstant(toMachineWord closureRef), dest=closureReg, kind=movePolyWord}) :: BlockSimple(AllocateMemoryOperation{size=n, flags=F_closure, dest=memAddr, saveRegs=[]}) :: tlCode end | loadFields(f :: rest, n, tlCode) = let (* Defer the evaluation if possible. We may have a constant that we can't move directly but it's better to load it after the allocation otherwise we will have to push the register if we need to GC. *) val (code1, source1, _) = codeToICodeRev(BICExtract f, context, false, Allowed allowDefer, tlCode) val restAndAlloc = loadFields(rest, n+1, code1) val (code2, source, _) = moveIfNotAllowedRev(Allowed allowInMemMove, restAndAlloc, source1) val storeValue = BlockSimple(StoreArgument{ source=source, offset=n*Word.toInt wordSize, base=memAddr, index=ObjectIndex, kind=movePolyWord, isMutable=false}) in storeValue :: code2 end val code = BlockSimple InitialisationComplete :: BlockSimple(LoadArgument{source=RegisterArgument memAddr, dest=target, kind=movePolyWord}) :: loadFields(closure, 2, tailCode) in (code, RegisterArgument target, false) end (* Treat it as a tuple with the code as the first field. *) else codeToICodeRev(BICTuple(BICConstnt(codeAddressFromClosure closureRef, []) :: map BICExtract closure), context, isTail, destination, tailCode) end | codeToICodeRev(BICCond(test, thenPt, elsePt), context, isTail, NoResult, tailCode) = let (* If we don't want the result but are only evaluating for side-effects we may be able to optimise special cases. This was easier in the forward case but for now we don't bother and leave it to the lower levels. *) val startElse = newLabel() and skipElse = newLabel() val codeTest = codeConditionRev(test, context, false, startElse, tailCode) val (codeThen, _, _) = codeToICodeRev(thenPt, context, isTail, NoResult, codeTest) val (codeElse, _, _) = codeToICodeRev(elsePt, context, isTail, NoResult, BlockLabel startElse :: BlockFlow(Unconditional skipElse) :: codeThen) in (BlockLabel skipElse :: codeElse, (* Unit result *) IntegerConstant(tag 0), false) end | codeToICodeRev(BICCond(test, thenPt, elsePt), context, isTail, destination, tailCode) = let (* Because we may push the result onto the stack we have to create a new preg to hold the result and then copy that to the final result. *) (* If this is a tail each arm will exit separately and neither will return a result. *) val target = asTarget destination val condResult = newMergeReg() val thenTarget = if isTail then newPReg() else condResult val startElse = newLabel() val testCode = codeConditionRev(test, context, false, startElse, tailCode) (* Put the result in the target register. *) val (thenCode, _, thenExited) = codeToICodeRev(thenPt, context, isTail, SpecificPReg thenTarget, testCode) (* Add a jump round the else-part except that if this is a tail we return. The then-part could have exited e.g. with a raise or a loop. *) val (exitThen, thenLabel, elseTarget) = if thenExited then (thenCode, [], target (* Can use original target. *)) else if isTail then (returnInstruction(context, thenTarget, thenCode), [], newPReg()) else let val skipElse = newLabel() in (BlockFlow(Unconditional skipElse) :: thenCode, [BlockSimple(LoadArgument{source=RegisterArgument condResult, dest=target, kind=movePolyWord}), BlockLabel skipElse], condResult) end val (elseCode, _, elseExited) = codeToICodeRev(elsePt, context, isTail, SpecificPReg elseTarget, BlockLabel startElse :: exitThen) (* Add a return to the else-part if necessary so we will always exit on a tail. *) val exitElse = if isTail andalso not elseExited then returnInstruction(context, elseTarget, elseCode) else elseCode in (thenLabel @ exitElse, RegisterArgument target, isTail orelse thenExited andalso elseExited) end | codeToICodeRev(BICCase { cases, test, default, isExhaustive, firstIndex}, context, isTail, destination, tailCode) = let (* We have to create a new preg for the result in case we need to push it to the stack. *) val targetReg = newMergeReg() local val initialTestReg = newPReg() val (testCode, _, _) = codeToICodeRev(test, context, false, SpecificPReg initialTestReg, tailCode) (* Subtract the minimum value so the value we're testing is always in the range of (tagged) 0 to the maximum. It is possible to adjust the value when computing the index but that can lead to overflows during compilation if the minimum is very large or small. We can ignore overflow and allow values to wrap round. *) in val (testCode, testReg) = if firstIndex = 0w0 then (testCode, initialTestReg) else let val newTestReg = newPReg() val subtract = BlockSimple(ArithmeticFunction{oper=SUB, resultReg=newTestReg, operand1=initialTestReg, operand2=IntegerConstant(semitag(Word.toLargeInt firstIndex)), ccRef=newCCRef(), opSize=polyWordOpSize}) in (subtract :: testCode, newTestReg) end end val workReg = newPReg() (* Unless this is exhaustive we need to add a range check. *) val (rangeCheck, extraDefaults) = if isExhaustive then (testCode, []) else let val defLab1 = newLabel() val tReg1 = newPReg() val ccRef1 = newCCRef() (* Since we've subtracted any minimum we only have to check whether the value is greater (unsigned) than the maximum. *) val numberOfCases = LargeInt.fromInt(List.length cases) val continueLab = newLabel() val testCode2 = BlockLabel continueLab :: BlockFlow(Conditional{ccRef=ccRef1, condition=JNB, trueJump=defLab1, falseJump=continueLab}) :: BlockSimple(WordComparison{arg1=tReg1, arg2=IntegerConstant(tag numberOfCases), ccRef=ccRef1, opSize=polyWordOpSize}) :: BlockSimple(LoadArgument {source=RegisterArgument testReg, dest=tReg1, kind=movePolyWord}) :: testCode in (testCode2, [defLab1]) end (* Make a label for each item in the list. *) val codeLabels = map (fn _ => newLabel()) cases (* Create an exit label in case it's needed. *) val labelForExit = if isTail then ~1 (* Illegal label. *) else newLabel() (* Generate the code for each of the cases and the default. We need to put an unconditional branch after each to skip the other cases. *) fun codeCases (SOME c :: otherCases, startLabel :: otherLabels, tailCode) = let val caseTarget = if isTail then newPReg() else targetReg (* Put in the case with a jump to the end of the sequence. *) val (codeThisCase, _, caseExited) = codeToICodeRev(c, context, isTail, SpecificPReg caseTarget, BlockLabel startLabel :: tailCode) val exitThisCase = if caseExited then codeThisCase else if isTail then returnInstruction(context, caseTarget, codeThisCase) else BlockFlow(Unconditional labelForExit) :: codeThisCase in codeCases(otherCases, otherLabels, exitThisCase) end | codeCases(NONE :: otherCases, _ :: otherLabels, tailCode) = codeCases(otherCases, otherLabels, tailCode) | codeCases ([], [], tailCode) = let (* We need to add labels for all the gaps we filled and also for a "default" label for the indexed-case instruction itself as well as any range checks. *) fun addDefault (startLabel, NONE, l) = BlockLabel startLabel :: l | addDefault (_, SOME _, l) = l fun asForward l = BlockLabel l val dLabs = map asForward extraDefaults @ tailCode val defLabels = ListPair.foldlEq addDefault dLabs (codeLabels, cases) val defaultTarget = if isTail then newPReg() else targetReg val (defaultCode, _, defaultExited) = codeToICodeRev(default, context, isTail, SpecificPReg defaultTarget, defLabels) in (* Put in the default. Because this is the last we don't need to jump round it. However if this is a tail and we haven't exited we put in a return. That way the case will always have exited if this is a tail. *) if isTail andalso not defaultExited then returnInstruction(context, defaultTarget, defaultCode) else defaultCode end | codeCases _ = raise InternalError "codeCases: mismatch" val codedCases = codeCases(cases, codeLabels, BlockFlow(IndexedBr codeLabels) :: BlockSimple(IndexedCaseOperation{testReg=testReg, workReg=workReg}) :: rangeCheck) (* We can now copy to the target. If we need to push the result this load will be converted into a push. *) val target = asTarget destination val copyToTarget = if isTail then codedCases else BlockSimple(LoadArgument{source=RegisterArgument targetReg, dest=target, kind=movePolyWord}) :: BlockLabel labelForExit :: codedCases in (copyToTarget, RegisterArgument target, isTail (* We have always exited on a tail. *)) end | codeToICodeRev(BICBeginLoop {loop, arguments}, context as { stackPtr, currHandler, overflowBlock, ...}, isTail, destination, tailCode) = let val target = asTarget destination fun codeArgs ([], tailCode) = ([], tailCode) | codeArgs (({value, addr}, _) :: rest, tailCode) = let val pr = newPReg() val () = Array.update(locToPregArray, addr, PregLocation pr) val (code, _, _) = codeToICodeRev(value, context, false, SpecificPReg pr, tailCode) val (pregs, othercode) = codeArgs(rest, code) in (pr::pregs, othercode) end val (loopRegs, argCode) = codeArgs(arguments, tailCode) val loopLabel = newLabel() val (loopBody, _, loopExited) = codeToICodeRev(loop, {loopArgs=SOME (loopRegs, loopLabel, stackPtr), stackPtr=stackPtr, currHandler=currHandler, overflowBlock=overflowBlock }, isTail, SpecificPReg target, BlockLabel loopLabel :: BlockSimple BeginLoop :: argCode) in (loopBody, RegisterArgument target, loopExited) end | codeToICodeRev(BICLoop args, context as {loopArgs=SOME (loopRegs, loopLabel, loopSp), stackPtr, currHandler, ...}, _, destination, tailCode) = let val target = asTarget destination (* Registers to receive the evaluated arguments. We can't put the values into the loop variables yet because the values could depend on the current values of the loop variables. *) val argPRegs = map(fn _ => newPReg()) args val codeArgs = ListPair.foldlEq(fn ((arg, _), pr, l) => #1 (codeToICodeRev(arg, context, false, SpecificPReg pr, l))) tailCode (args, argPRegs) val jumpArgs = ListPair.mapEq(fn (s, l) => (RegisterArgument s, l)) (argPRegs, loopRegs) (* If we've allocated a container in the loop we have to remove it before jumping back. *) val stackReset = if loopSp = stackPtr then codeArgs else BlockSimple(ResetStackPtr{numWords=stackPtr-loopSp, preserveCC=false}) :: codeArgs val jumpLoop = JumpLoop{regArgs=jumpArgs, stackArgs=[], checkInterrupt=SOME[], workReg=NONE} (* "checkInterrupt" could result in a Interrupt exception so we treat this like a function call. *) val code = case currHandler of NONE => BlockFlow(Unconditional loopLabel) :: BlockSimple jumpLoop :: stackReset | SOME h => BlockOptionalHandle{call=jumpLoop, handler=h, label=loopLabel} :: stackReset in (code, RegisterArgument target, true) end | codeToICodeRev(BICLoop _, {loopArgs=NONE, ...}, _, _, _) = raise InternalError "BICLoop without BICBeginLoop" | codeToICodeRev(BICRaise exc, context as { currHandler, ...}, _, destination, tailCode) = let val packetReg = newPReg() val (code, _, _) = codeToICodeRev(exc, context, false, SpecificPReg packetReg, tailCode) val raiseCode = RaiseExceptionPacket{packetReg=packetReg} val block = case currHandler of NONE => BlockExit raiseCode | SOME h => BlockRaiseAndHandle(raiseCode, h) in (block :: code, RegisterArgument(asTarget destination), true (* Always exits *)) end | codeToICodeRev(BICHandle{exp, handler, exPacketAddr}, context as { stackPtr, loopArgs, overflowBlock, ... }, isTail, destination, tailCode) = let (* As with BICCond and BICCase we need to create a new register for the result in case we need to push it to the stack. *) val handleResult = newMergeReg() val handlerLab = newLabel() and startHandling = newLabel() val (bodyTarget, handlerTarget) = if isTail then (newPReg(), newPReg()) else (handleResult, handleResult) (* TODO: Even if we don't actually want a result we force one in here by using "asTarget". *) (* The expression cannot be treated as a tail because the handler has to be removed after. It may "exit" if it has raised an unconditional exception. If it has we mustn't generate a PopExceptionHandler because there won't be any result for resultReg. We need to add two words to the stack to account for the items pushed by PushExceptionHandler. We create an instruction to push the handler followed by a block fork to the start of the code and, potentially the handler, then a label to start the code that the handler is in effect for. *) val initialCode = BlockLabel startHandling :: BlockFlow(SetHandler{handler=handlerLab, continue=startHandling}) :: BlockSimple(PushExceptionHandler{workReg=newPReg()}) :: tailCode val (expCode, _, expExit) = codeToICodeRev(exp, {stackPtr=stackPtr+2, loopArgs=loopArgs, currHandler=SOME handlerLab, overflowBlock=overflowBlock}, false (* Not tail *), SpecificPReg bodyTarget, initialCode) (* If this is the tail we can replace the jump at the end of the handled code with returns. If the handler has exited we don't need a return there. Otherwise we need to add an unconditional jump to skip the handler. *) val (atExpEnd, skipExpLabel) = case (isTail, expExit) of (true, true) => (* Tail and exited. *) (expCode, NONE) | (true, false) => (* Tail and not exited. *) (returnInstruction(context, bodyTarget, BlockSimple(PopExceptionHandler{workReg=newPReg()}) :: expCode), NONE) | (false, true) => (* Not tail but exited. *) (expCode, NONE) | (false, false) => let val skipHandler = newLabel() in (BlockFlow(Unconditional skipHandler) :: BlockSimple(PopExceptionHandler{workReg=newPReg()}) :: expCode, SOME skipHandler) end (* Make a register to hold the exception packet and put eax into it. *) val packetAddr = newPReg() val () = Array.update(locToPregArray, exPacketAddr, PregLocation packetAddr) val (handleCode, _, handleExit) = codeToICodeRev(handler, context, isTail, SpecificPReg handlerTarget, BlockSimple(BeginHandler{workReg=newPReg(), packetReg=packetAddr}) :: BlockLabel handlerLab :: atExpEnd) val target = asTarget destination val afterHandler = case (isTail, handleExit) of (true, true) => (* Tail and exited. *) handleCode | (true, false) => (* Tail and not exited. *) returnInstruction(context, handlerTarget, handleCode) | (false, _) => (* Not tail. *) handleCode val addLabel = case skipExpLabel of SOME lab => BlockLabel lab:: afterHandler | NONE => afterHandler in (BlockSimple(LoadArgument{source=RegisterArgument handleResult, dest=target, kind=movePolyWord}) :: addLabel, RegisterArgument target, isTail) end | codeToICodeRev(BICTuple fields, context, _, destination, tailCode) = let (* TODO: This is a relic of the old fall-back code-generator. It required the result of a tuple to be at the top of the stack. It should be changed. *) val target = asTarget destination (* Actually we want this. *) val memAddr = newPReg() fun loadFields([], n, tlCode) = BlockSimple(AllocateMemoryOperation{size=n, flags=0w0, dest=memAddr, saveRegs=[]}) :: tlCode | loadFields(f :: rest, n, tlCode) = let (* Defer the evaluation if possible. We may have a constant that we can't move directly but it's better to load it after the allocation otherwise we will have to push the register if we need to GC. *) val (code1, source1, _) = codeToICodeRev(f, context, false, Allowed allowDefer, tlCode) val restAndAlloc = loadFields(rest, n+1, code1) val (code2, source, _) = moveIfNotAllowedRev(Allowed allowInMemMove, restAndAlloc, source1) val storeValue = BlockSimple(StoreArgument{ source=source, offset=n*Word.toInt wordSize, base=memAddr, index=memIndexOrObject, kind=movePolyWord, isMutable=false}) in storeValue :: code2 end val code = BlockSimple InitialisationComplete :: BlockSimple(LoadArgument{source=RegisterArgument memAddr, dest=target, kind=movePolyWord}) :: loadFields(fields, 0, tailCode) in (code, RegisterArgument target, false) end (* Copy the source tuple into the container. There are important special cases for both the source tuple and the container. If the source tuple is a BICTuple we have the fields and can store them without creating a tuple on the heap. If the destination is a local container we can store directly into the stack. *) | codeToICodeRev(BICSetContainer{container, tuple, filter}, context as {stackPtr, ...}, _, destination, tailCode) = let local fun createStore containerReg (source, destWord) = StoreArgument{source=source, offset=destWord*Word.toInt nativeWordSize, base=containerReg, index=NoMemIndex, kind=moveNativeWord, isMutable=false} in val findContainer = case container of BICExtract(BICLoadLocal l) => ( case Array.sub(locToPregArray, l) of ContainerLocation{container, stackOffset} => let fun storeToStack(source, destWord) = StoreToStack{source=source, container=container, field=destWord, stackOffset=stackPtr-stackOffset+destWord} in SOME storeToStack end | _ => NONE ) | _ => NONE val (codeContainer, storeInstr) = case findContainer of SOME storeToStack => (tailCode, storeToStack) | NONE => let val containerTarget = newPReg() val (codeContainer, _, _) = codeToICodeRev(container, context, false, SpecificPReg containerTarget, tailCode) in (codeContainer, createStore containerTarget) end end val filterLength = BoolVector.length filter val code = case tuple of BICTuple cl => let (* In theory it's possible that the tuple could contain fields that are not used but nevertheless need to be evaluated for their side-effects. Create all the fields and push to the stack. *) fun codeField(arg, (regs, tailCode)) = let val (c, r, _) = codeToICodeRev(arg, context, false, Allowed allowInMemMove, tailCode) in (r :: regs, c) end val (pregsRev, codeFields) = List.foldl codeField ([], codeContainer) cl val pregs = List.rev pregsRev fun copyField(srcReg, (sourceWord, destWord, tailCode)) = if sourceWord < filterLength andalso BoolVector.sub(filter, sourceWord) then (sourceWord+1, destWord+1, BlockSimple(storeInstr(srcReg, destWord)) :: tailCode) else (sourceWord+1, destWord, tailCode) val (_, _, resultCode) = List.foldl copyField (0, 0, codeFields) pregs in resultCode end | tuple => let (* Copy a heap tuple. It is possible that this is another container in which case we must load the fields directly. We mustn't load its address and then copy because loading the address would be the last reference and might cause the container to be reused prematurely. *) val findContainer = case tuple of BICExtract(BICLoadLocal l) => ( case Array.sub(locToPregArray, l) of ContainerLocation{container, stackOffset} => let fun getAddr sourceWord = StackLocation{wordOffset=stackPtr-stackOffset+sourceWord, container=container, field=sourceWord, cache=NONE} in SOME getAddr end | _ => NONE ) | _ => NONE val (codeTuple, loadField) = case findContainer of SOME getAddr => (codeContainer, getAddr) | NONE => let val tupleTarget = newPReg() val (codeTuple, _, _) = codeToICodeRev(tuple, context, false, SpecificPReg tupleTarget, codeContainer) fun loadField sourceWord = wordOffsetAddress(sourceWord, tupleTarget) in (codeTuple, loadField) end fun copyContainer(sourceWord, destWord, tailCode) = if sourceWord = filterLength then tailCode else if BoolVector.sub(filter, sourceWord) then let val loadReg = newPReg() val code = BlockSimple(storeInstr(RegisterArgument loadReg, destWord)) :: BlockSimple(LoadArgument{source=loadField sourceWord, dest=loadReg, kind=movePolyWord}) :: tailCode in copyContainer(sourceWord+1, destWord+1, code) end else copyContainer(sourceWord+1, destWord, tailCode) in copyContainer(0, 0, codeTuple) end in moveIfNotAllowedRev(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeRev(BICTagTest{test, tag=tagValue, ...}, context, _, destination, tailCode) = (* Check the "tag" word of a union (datatype). N.B. Not the same as testing the tag bit of a word. *) let val ccRef = newCCRef() val memOrReg = { anyConstant=false, const32s=false, memAddr=true, existingPreg=true } val (testCode, tagArg, _) = codeToICodeRev(test, context, false, Allowed memOrReg, tailCode) val target = asTarget destination in (makeBoolResultRev(JE, ccRef, target, (* Use CompareLiteral because the tag must fit in 32-bits. *) BlockSimple(CompareLiteral{arg1=tagArg, arg2=tag(Word.toLargeInt tagValue), opSize=polyWordOpSize, ccRef=ccRef}) :: testCode), RegisterArgument target, false) end | codeToICodeRev(BICLoadOperation instr, context, isTail, destination, tailCode) = let val (code, dest, haveExited) = codeToICodeLoad(instr, context, isTail, destination) in (revApp(code, tailCode), dest, haveExited) end | codeToICodeRev(BICStoreOperation instr, context, isTail, destination, tailCode) = let val (code, dest, haveExited) = codeToICodeStore(instr, context, isTail, destination) in (revApp(code, tailCode), dest, haveExited) end | codeToICodeRev(BICBlockOperation ({kind=BlockOpEqualByte, sourceLeft, destRight, length}), context, _, destination, tailCode) = let val vec1Reg = newUReg() and vec2Reg = newUReg() val ccRef = newCCRef() val (leftCode, leftUntag, {base=leftBase, offset=leftOffset, index=leftIndex, ...}) = codeAddressRev(sourceLeft, true, context, tailCode) val (rightCode, rightUntag, {base=rightBase, offset=rightOffset, index=rightIndex, ...}) = codeAddressRev(destRight, true, context, leftCode) val (lengthCode, lengthUntag, lengthArg) = codeAsUntaggedToRegRev(length, false (* unsigned *), context, rightCode) val target = asTarget destination val code = makeBoolResultRev(JE, ccRef, target, BlockSimple(CompareByteVectors{ vec1Addr=vec1Reg, vec2Addr=vec2Reg, length=lengthArg, ccRef=ccRef }) :: lengthUntag @ BlockSimple(loadAddress{base=rightBase, offset=rightOffset, index=rightIndex, dest=vec2Reg}) :: rightUntag @ BlockSimple(loadAddress{base=leftBase, offset=leftOffset, index=leftIndex, dest=vec1Reg}) :: leftUntag @ lengthCode) in (code, RegisterArgument target, false) end | codeToICodeRev(BICBlockOperation instr, context, isTail, destination, tailCode) = let val (code, dest, haveExited) = codeToICodeBlock(instr, context, isTail, destination) in (revApp(code, tailCode), dest, haveExited) end and codeToICodeUnaryRev({oper=BuiltIns.NotBoolean, arg1}, context, _, destination, tailCode) = let val target = asTarget destination val ccRef = newCCRef() val allow = Allowed {anyConstant=false, const32s=false, memAddr=true, existingPreg=true} val (argCode, testDest, _) = codeToICodeRev(arg1, context, false, allow, tailCode) in (* Test the argument and return a boolean result. If either the argument is a condition or the result is used in a test this will be better than using XOR. *) (makeBoolResultRev(JNE, ccRef, target, BlockSimple(CompareLiteral{arg1=testDest, arg2=tag 1, opSize=polyWordOpSize, ccRef=ccRef}) :: argCode), RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.IsTaggedValue, arg1}, context, _, destination, tailCode) = let val ccRef = newCCRef() val memOrReg = { anyConstant=false, const32s=false, memAddr=true, existingPreg=true } val (testCode, testResult, _) = codeToICodeRev(arg1, context, false, Allowed memOrReg, tailCode) (* Test the tag bit. This sets the zero bit if the value is untagged. *) val target = asTarget destination in (makeBoolResultRev(JNE, ccRef, target, BlockSimple(TestTagBit{arg=testResult, ccRef=ccRef}) :: testCode), RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.MemoryCellLength, arg1}, context, _, destination, tailCode) = let val target = asTarget destination val argReg1 = newUReg() and argReg2 = newUReg() and argReg3 = newUReg() (* These are untagged until the tag is put in. *) and ccRef1 = newCCRef() and ccRef2 = newCCRef() and ccRef3 = newCCRef() (* Get the length of a memory cell (heap object). We need to mask out the top byte containing the flags and to tag the result. The mask is 56 bits on 64-bit which won't fit in an inline constant. Since we have to shift it anyway we might as well do this by shifts. *) val (argCode, addrReg) = codeToPRegRev(arg1, context, tailCode) in (BlockSimple(ArithmeticFunction{oper=OR, resultReg=target, operand1=argReg3, operand2=IntegerConstant 1, ccRef=ccRef3, opSize=polyWordOpSize}) :: BlockSimple(ShiftOperation{shift=SHR, resultReg=argReg3, operand=argReg2, shiftAmount=IntegerConstant 7 (* 8-tagshift*), ccRef=ccRef2, opSize=polyWordOpSize }) :: BlockSimple(ShiftOperation{shift=SHL, resultReg=argReg2, operand=argReg1, shiftAmount=IntegerConstant 8, ccRef=ccRef1, opSize=polyWordOpSize }) :: BlockSimple(LoadArgument{source=wordOffsetAddress(~1, addrReg), dest=argReg1, kind=movePolyWord}) :: argCode, RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.MemoryCellFlags, arg1}, context, _, destination, tailCode) = let val target = asTarget destination val argReg1 = newUReg() val (argCode, addrReg) = codeToPRegRev(arg1, context, tailCode) in (BlockSimple(TagValue{ source=argReg1, dest=target, isSigned=false, opSize=OpSize32 }) :: BlockSimple(LoadArgument{source=MemoryLocation{offset= ~1, base=addrReg, index=memIndexOrObject, cache=NONE}, dest=argReg1, kind=MoveByte}) :: argCode, RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.ClearMutableFlag, arg1}, context, _, destination, tailCode) = let val (argCode, addrReg) = codeToPRegRev(arg1, context, tailCode) in moveIfNotAllowedRev(destination, BlockSimple(LockMutable{addr=addrReg}) :: argCode, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeUnaryRev({oper=BuiltIns.AtomicIncrement, arg1}, context, _, destination, tailCode) = let val target = asTarget destination val incrReg = newUReg() val (argCode, addrReg) = codeToPRegRev(arg1, context, tailCode) val code = (* We want the result to be the new value but we've returned the old value. *) BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=incrReg, operand2=IntegerConstant(semitag 1), ccRef=newCCRef(), opSize=polyWordOpSize}) :: BlockSimple(AtomicExchangeAndAdd{ base=addrReg, source=incrReg }) :: BlockSimple(LoadArgument{source=IntegerConstant(semitag 1), dest=incrReg, kind=movePolyWord}) :: argCode in (code, RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.AtomicDecrement, arg1}, context, _, destination, tailCode) = let val target = asTarget destination val incrReg = newUReg() val (argCode, addrReg) = codeToPRegRev(arg1, context, tailCode) val code = BlockSimple(ArithmeticFunction{oper=SUB, resultReg=target, operand1=incrReg, operand2=IntegerConstant(semitag 1), ccRef=newCCRef(), opSize=polyWordOpSize}) :: BlockSimple(AtomicExchangeAndAdd{ base=addrReg, source=incrReg }) :: BlockSimple(LoadArgument{source=IntegerConstant(semitag ~1), dest=incrReg, kind=movePolyWord}) :: argCode in (code, RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.AtomicReset, arg1}, context, _, destination, tailCode) = let (* This is needed only for the interpreted version where we have a single real mutex to interlock atomic increment and decrement. We have to use the same mutex to interlock clearing a mutex. On the X86 we use hardware locking and the hardware guarantees that an assignment of a word will be atomic. *) val (argCode, addrReg) = codeToPRegRev(arg1, context, tailCode) (* Store tagged 1 in the mutex. This is the unlocked value. *) val code = BlockSimple(StoreArgument{source=IntegerConstant(tag 1), base=addrReg, index=memIndexOrObject, offset=0, kind=movePolyWord, isMutable=true}) :: argCode in moveIfNotAllowedRev(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeUnaryRev({oper=BuiltIns.LongWordToTagged, arg1}, context, _, destination, tailCode) = let (* This is exactly the same as StringLengthWord at the moment. TODO: introduce a new ICode entry so that the next stage can optimise longword operations. *) val target = asTarget destination val argReg1 = newUReg() val (argCode, addrReg) = codeToPRegRev(arg1, context, tailCode) val code = BlockSimple(TagValue{ source=argReg1, dest=target, isSigned=false, opSize=polyWordOpSize }) :: BlockSimple(LoadArgument{source=wordAt addrReg, dest=argReg1, kind=movePolyWord}) :: argCode in (code, RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.SignedToLongWord, arg1}, context, _, destination, tailCode) = let val addrReg = newPReg() and untagArg = newUReg() val (argCode, argReg1) = codeToPRegRev(arg1, context, tailCode) val (signExtend, sxReg) = case targetArch of ObjectId32Bit => let val sReg = newUReg() in ([BlockSimple(SignExtend32To64{source=RegisterArgument argReg1, dest=sReg})], sReg) end | _ => ([], argReg1) val code = BlockSimple(BoxValue{boxKind=BoxLargeWord, source=untagArg, dest=addrReg, saveRegs=[]}) :: BlockSimple(UntagValue{source=sxReg, dest=untagArg, isSigned=true, cache=NONE, opSize=nativeWordOpSize}) :: signExtend @ argCode in moveIfNotAllowedRev(destination, code, RegisterArgument addrReg) end | codeToICodeUnaryRev({oper=BuiltIns.UnsignedToLongWord, arg1}, context, _, destination, tailCode) = let val addrReg = newPReg() and untagArg = newUReg() val (argCode, argReg1) = codeToPRegRev(arg1, context, tailCode) val code = BlockSimple(BoxValue{boxKind=BoxLargeWord, source=untagArg, dest=addrReg, saveRegs=[]}) :: (* We can just use a polyWord operation to untag the unsigned value. *) BlockSimple(UntagValue{source=argReg1, dest=untagArg, isSigned=false, cache=NONE, opSize=polyWordOpSize}) :: argCode in moveIfNotAllowedRev(destination, code, RegisterArgument addrReg) end | codeToICodeUnaryRev({oper=BuiltIns.RealNeg precision, arg1}, context, _, destination, tailCode) = let val target = asTarget destination val fpRegSrc = newUReg() and fpRegDest = newUReg() and sse2ConstReg = newUReg() (* The SSE2 code uses an SSE2 logical operation to flip the sign bit. This requires the values to be loaded into registers first because the logical operations require 128-bit operands. *) val (argCode, aReg1) = codeToPReg(arg1, context) (* Double precision values are always boxed and single precision values if they won't fit in a word. Otherwise we can using tagging. *) open BuiltIns val load = if precision = PrecDouble then BlockSimple(LoadArgument{source=wordAt aReg1, dest=fpRegSrc, kind=MoveDouble}) else if wordSize = 0w8 then BlockSimple(UntagFloat{source=RegisterArgument aReg1, dest=fpRegSrc, cache=NONE}) else BlockSimple(LoadArgument{source=wordAt aReg1, dest=fpRegSrc, kind=MoveFloat}) val code = case fpMode of FPModeX87 => [BlockSimple(X87FPUnaryOps{ fpOp=FCHS, dest=fpRegDest, source=fpRegSrc})] | FPModeSSE2 => let (* In single precision mode the sign bit is in the low 32-bits. There may be a better way to load it. *) val signBit = if precision = PrecDouble then realSignBit else floatSignBit in [BlockSimple(LoadArgument{source=AddressConstant signBit, dest=sse2ConstReg, kind=MoveDouble}), BlockSimple(SSE2FPBinary{opc=SSE2BXor, resultReg=fpRegDest, arg1=fpRegSrc, arg2=RegisterArgument sse2ConstReg})] end val result = boxOrTagReal(fpRegDest, target, precision) in (revApp(argCode @ load :: code @ result, tailCode), RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.RealAbs precision, arg1}, context, _, destination, tailCode) = let val target = asTarget destination val fpRegSrc = newUReg() and fpRegDest = newUReg() and sse2ConstReg = newUReg() val (argCode, aReg1) = codeToPReg(arg1, context) open BuiltIns val load = if precision = PrecDouble then BlockSimple(LoadArgument{source=wordAt aReg1, dest=fpRegSrc, kind=MoveDouble}) else if wordSize = 0w8 then BlockSimple(UntagFloat{source=RegisterArgument aReg1, dest=fpRegSrc, cache=NONE}) else BlockSimple(LoadArgument{source=wordAt aReg1, dest=fpRegSrc, kind=MoveFloat}) val code = case fpMode of FPModeX87 => [BlockSimple(X87FPUnaryOps{ fpOp=FABS, dest=fpRegDest, source=fpRegSrc})] | FPModeSSE2 => let val mask = if precision = PrecDouble then realAbsMask else floatAbsMask in [BlockSimple(LoadArgument{source=AddressConstant mask, dest=sse2ConstReg, kind=MoveDouble}), BlockSimple(SSE2FPBinary{opc=SSE2BAnd, resultReg=fpRegDest, arg1=fpRegSrc, arg2=RegisterArgument sse2ConstReg})] end val result = boxOrTagReal(fpRegDest, target, precision) in (revApp(argCode @ load :: code @ result, tailCode), RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.RealFixedInt precision, arg1}, context, _, destination, tailCode) = let val target = asTarget destination val untagReg = newUReg() and fpReg = newUReg() val (argCode, aReg1) = codeToPReg(arg1, context) val floatOp = case fpMode of FPModeX87 => X87Float | FPModeSSE2 => SSE2Float val boxFloat = case fpMode of FPModeX87 => BoxX87Double | FPModeSSE2 => BoxSSE2Double val _ = precision = BuiltIns.PrecDouble orelse raise InternalError "RealFixedInt - single" val code = argCode @ [BlockSimple(UntagValue{source=aReg1, dest=untagReg, isSigned=true, cache=NONE, opSize=polyWordOpSize}), BlockSimple(floatOp{ dest=fpReg, source=RegisterArgument untagReg}), BlockSimple(BoxValue{boxKind=boxFloat, source=fpReg, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.FloatToDouble, arg1}, context, _, destination, tailCode) = let (* Convert a single precision floating point value to double precision. *) val target = asTarget destination val fpReg = newUReg() and fpReg2 = newUReg() val (argCode, aReg1) = codeToPReg(arg1, context) (* MoveFloat always converts from single to double-precision. *) val unboxOrUntag = case (fpMode, wordSize) of (FPModeX87, _) => [BlockSimple(LoadArgument{source=wordAt aReg1, dest=fpReg2, kind=MoveFloat})] | (FPModeSSE2, 0w4) => [BlockSimple(LoadArgument{source=wordAt aReg1, dest=fpReg, kind=MoveFloat}), BlockSimple(SSE2FPUnary{opc=SSE2UFloatToDouble, resultReg=fpReg2, source=RegisterArgument fpReg})] | (FPModeSSE2, _) => [BlockSimple(UntagFloat{source=RegisterArgument aReg1, dest=fpReg, cache=NONE}), BlockSimple(SSE2FPUnary{opc=SSE2UFloatToDouble, resultReg=fpReg2, source=RegisterArgument fpReg})] val boxFloat = case fpMode of FPModeX87 => BoxX87Double | FPModeSSE2 => BoxSSE2Double val code = argCode @ unboxOrUntag @ [BlockSimple(BoxValue{boxKind=boxFloat, source=fpReg2, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.DoubleToFloat NONE, arg1}, context, _, destination, tailCode) = let (* Convert a double precision value to a single precision using the current rounding mode. This is simpler than setting the rounding mode and then restoring it. *) val target = asTarget destination val fpReg = newUReg() and fpReg2 = newUReg() val (argCode, aReg1) = codeToPReg(arg1, context) (* In 32-bit mode we need to box the float. In 64-bit mode we can tag it. *) val boxOrTag = case fpMode of FPModeX87 => [BlockSimple(BoxValue{boxKind=BoxX87Float, source=fpReg, dest=target, saveRegs=[]})] | FPModeSSE2 => BlockSimple(SSE2FPUnary{opc=SSE2UDoubleToFloat, resultReg=fpReg2, source=RegisterArgument fpReg}) :: boxOrTagReal(fpReg2, target, BuiltIns.PrecSingle) val code = argCode @ [BlockSimple(LoadArgument{source=wordAt aReg1, dest=fpReg, kind=MoveDouble})] @ boxOrTag in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.DoubleToFloat (SOME rndMode), arg1}, context, _, destination, tailCode) = let (* Convert a double precision value to a single precision. The rounding mode is passed in explicitly. *) val target = asTarget destination val fpReg = newUReg() and fpReg2 = newUReg() val (argCode, aReg1) = codeToPReg(arg1, context) (* In 32-bit mode we need to box the float. In 64-bit mode we can tag it. *) (* We need to save the rounding mode before we change it and restore it afterwards. *) open IEEEReal fun doConversion() = case fpMode of FPModeX87 => (* Convert the value using the appropriate rounding. *) [BlockSimple(BoxValue{boxKind=BoxX87Float, source=fpReg, dest=target, saveRegs=[]})] | FPModeSSE2 => BlockSimple(SSE2FPUnary{opc=SSE2UDoubleToFloat, resultReg=fpReg2, source=RegisterArgument fpReg}) :: boxOrTagReal(fpReg2, target, BuiltIns.PrecSingle) val code = argCode @ [BlockSimple(LoadArgument{source=wordAt aReg1, dest=fpReg, kind=MoveDouble})] @ setAndRestoreRounding(rndMode, doConversion) in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.RealToInt(precision, rndMode), arg1}, context, _, destination, tailCode) = let val target = asTarget destination val chkOverflow = newCCRef() val convResult = newUReg() and wrkReg2 = newUReg() (* Convert a floating point value to an integer. We need to raise overflow if the result is out of range. We first convert the value to 32/64 bits then tag it. An overflow can happen either because the real number does not fit in 32/64 bits or if it is not a 31/63 bit value. Fortunately, if the first conversion fails the result is a value that causes an overflow when we try it shift it so the check for overflow only needs to happen there. There is an SSE2 instruction that implements truncation (round to zero) directly but in other cases we need to set the rounding mode. *) val doConvert = case (fpMode, precision) of (FPModeX87, _) => let val fpReg = newUReg() val (argCode, aReg) = codeToPReg(arg1, context) fun doConvert() = [BlockSimple(X87RealToInt{source=fpReg, dest=convResult })] in argCode @ [BlockSimple(LoadArgument{source=wordAt aReg, dest=fpReg, kind=MoveDouble})] @ setAndRestoreRounding(rndMode, doConvert) end | (FPModeSSE2, BuiltIns.PrecDouble) => let val (argCode, argReg) = codeToPReg(arg1, context) fun doConvert() = [BlockSimple( SSE2RealToInt{source=wordAt argReg, dest=convResult, isDouble=true, isTruncate = rndMode = IEEEReal.TO_ZERO }) ] in argCode @ ( case rndMode of IEEEReal.TO_ZERO => doConvert() | _ => setAndRestoreRounding(rndMode, doConvert)) end | (FPModeSSE2, BuiltIns.PrecSingle) => let val (argCode, aReg) = codeToPReg(arg1, context) val fpReg = newUReg() fun doConvert() = [BlockSimple( SSE2RealToInt{source=RegisterArgument fpReg, dest=convResult, isDouble=false, isTruncate = rndMode = IEEEReal.TO_ZERO })] in argCode @ [BlockSimple(UntagFloat{source=RegisterArgument aReg, dest=fpReg, cache=NONE})] @ ( case rndMode of IEEEReal.TO_ZERO => doConvert() | _ => setAndRestoreRounding(rndMode, doConvert) ) end val checkAndTag = BlockSimple(ShiftOperation{ shift=SHL, resultReg=wrkReg2, operand=convResult, shiftAmount=IntegerConstant 1, ccRef=chkOverflow, opSize=polyWordOpSize}) :: checkOverflow context chkOverflow @ [BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=wrkReg2, operand2=IntegerConstant 1, ccRef = newCCRef(), opSize=polyWordOpSize})] in (revApp(doConvert @ checkAndTag, tailCode), RegisterArgument target, false) end | codeToICodeUnaryRev({oper=BuiltIns.TouchAddress, arg1}, context, _, destination, tailCode) = let (* Put the value in a register. This is not entirely necessary but ensures that if the value is a constant the constant will be included in the code. *) val (argCode, aReg) = codeToPRegRev(arg1, context, tailCode) in moveIfNotAllowedRev(destination, BlockSimple(TouchArgument{source=aReg}) :: argCode, (* Unit result *) IntegerConstant(tag 0)) end - + and codeToICodeBinaryRev({oper=BuiltIns.WordComparison{test, isSigned}, arg1, arg2=BICConstnt(arg2Value, _)}, context, _, destination, tailCode) = let (* Comparisons. Because this is also used for pointer equality and even for exception matching it - is perfectly possible that the argument could be an address. *) + is perfectly possible that the argument could be an address. + The higher levels used to generate this for pointer equality. *) val ccRef = newCCRef() val comparison = (* If the argument is a tagged value that will fit in 32-bits we can use the literal version. Use toLargeIntX here because the value will be sign-extended even if we're actually doing an unsigned comparison. *) if isShort arg2Value andalso is32bit(tag(Word.toLargeIntX(toShort arg2Value))) then let val allow = Allowed {anyConstant=false, const32s=false, memAddr=true, existingPreg=true} in (* We're often comparing with a character or a string length field that has to be untagged. In that case we can avoid loading it into a register and untagging it by doing the comparison directly. *) case arg1 of BICLoadOperation{kind=LoadStoreUntaggedUnsigned, address} => let val (codeBaseIndex, codeUntag, memLoc) = codeAddressRev(address, false, context, tailCode) val literal = Word.toLargeIntX(toShort arg2Value) in BlockSimple(CompareLiteral{arg1=MemoryLocation memLoc, arg2=literal, opSize=polyWordOpSize, ccRef=ccRef}) :: codeUntag @ codeBaseIndex end | BICLoadOperation{kind=LoadStoreMLByte _, address} => let val (codeBaseIndex, codeUntag, {base, index, offset, ...}) = codeAddressRev(address, true, context, tailCode) val _ = toShort arg2Value >= 0w0 andalso toShort arg2Value < 0w256 orelse raise InternalError "Compare byte not a byte" val literal = Word8.fromLargeWord(Word.toLargeWord(toShort arg2Value)) in BlockSimple(CompareByteMem{arg1={base=base, index=index, offset=offset}, arg2=literal, ccRef=ccRef}) :: codeUntag @ codeBaseIndex end | BICUnary({oper=BuiltIns.MemoryCellFlags, arg1}) => (* This occurs particularly in arbitrary precision comparisons. *) let val (baseCode, baseReg) = codeToPRegRev(arg1, context, tailCode) val _ = toShort arg2Value >= 0w0 andalso toShort arg2Value < 0w256 orelse raise InternalError "Compare memory cell not a byte" val literal = Word8.fromLargeWord(Word.toLargeWord(toShort arg2Value)) in BlockSimple(CompareByteMem{arg1={base=baseReg, index=memIndexOrObject, offset= ~1}, arg2=literal, ccRef=ccRef}) :: baseCode end | _ => let (* TODO: We could include rarer cases of tagging by looking at the code and seeing if it's a TagValue. *) val (testCode, testDest, _) = codeToICodeRev(arg1, context, false, allow, tailCode) val literal = tag(Word.toLargeIntX(toShort arg2Value)) in BlockSimple(CompareLiteral{arg1=testDest, arg2=literal, opSize=polyWordOpSize, ccRef=ccRef}) :: testCode end end else (* Addresses or larger values. We need to use a register comparison. *) let val (testCode, testReg) = codeToPRegRev(arg1, context, tailCode) val arg2Arg = constantAsArgument arg2Value in BlockSimple(WordComparison{arg1=testReg, arg2=arg2Arg, ccRef=ccRef, opSize=polyWordOpSize}) :: testCode end val target = asTarget destination in (makeBoolResultRev(testAsBranch(test, isSigned, true), ccRef, target, comparison), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordComparison{test, isSigned}, arg1=BICConstnt(arg1Value, _), arg2}, context, _, destination, tailCode) = let (* If we have the constant first we need to reverse the test so the first argument is a register. *) val ccRef = newCCRef() val comparison = if isShort arg1Value andalso is32bit(tag(Word.toLargeIntX(toShort arg1Value))) then let val allow = Allowed {anyConstant=false, const32s=false, memAddr=true, existingPreg=true} val (testCode, testDest, _) = codeToICodeRev(arg2, context, false, allow, tailCode) val literal = tag(Word.toLargeIntX(toShort arg1Value)) in BlockSimple(CompareLiteral{arg1=testDest, arg2=literal, opSize=polyWordOpSize, ccRef=ccRef}) :: testCode end else (* Addresses or larger values. We need to use a register comparison. *) let val (testCode, testReg) = codeToPRegRev(arg2, context, tailCode) val arg1Arg = constantAsArgument arg1Value in BlockSimple(WordComparison{arg1=testReg, arg2=arg1Arg, ccRef=ccRef, opSize=polyWordOpSize}) :: testCode end val target = asTarget destination in (makeBoolResultRev(testAsBranch(leftRightTest test, isSigned, true), ccRef, target, comparison), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordComparison {test, isSigned}, arg1, arg2}, context, _, destination, tailCode) = let val ccRef = newCCRef() val memOrReg = { anyConstant=false, const32s=false, memAddr=true, existingPreg=true } val (arg1Code, arg1Result, _) = codeToICodeRev(arg1, context, false, Allowed memOrReg, tailCode) val (arg2Code, arg2Result, _) = codeToICodeRev(arg2, context, false, Allowed memOrReg, arg1Code) val target = asTarget destination val code = case (arg1Result, arg2Result) of (RegisterArgument arg1Reg, arg2Result) => makeBoolResultRev(testAsBranch(test, isSigned, true), ccRef, target, BlockSimple(WordComparison{arg1=arg1Reg, arg2=arg2Result, ccRef=ccRef, opSize=polyWordOpSize}) :: arg2Code) | (arg1Result, RegisterArgument arg2Reg) => (* The second argument is in a register - switch the sense of the test. *) makeBoolResultRev(testAsBranch(leftRightTest test, isSigned, true), ccRef, target, BlockSimple(WordComparison{arg1=arg2Reg, arg2=arg1Result, ccRef=ccRef, opSize=polyWordOpSize}) :: arg2Code) | (arg1Result, arg2Result) => let (* Have to load an argument - pick the first. *) val arg1Reg = newPReg() in makeBoolResultRev(testAsBranch(test, isSigned, true), ccRef, target, BlockSimple(WordComparison{arg1=arg1Reg, arg2=arg2Result, ccRef=ccRef, opSize=polyWordOpSize}) :: BlockSimple(LoadArgument{source=arg1Result, dest=arg1Reg, kind=movePolyWord}) :: arg2Code) end in (code, RegisterArgument target, false) end + | codeToICodeBinaryRev({oper=BuiltIns.PointerEq, arg1, arg2}, context, isTail, destination, tailCode) = + (* Equality of general values which can include pointers. This can be treated exactly as a word equality. + It has to be analysed differently for indexed cases. *) + codeToICodeBinaryRev({oper=BuiltIns.WordComparison{test=BuiltIns.TestEqual, isSigned=false}, arg1=arg1, arg2=arg2}, + context, isTail, destination, tailCode) + | codeToICodeBinaryRev({oper=BuiltIns.FixedPrecisionArith oper, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val code = codeFixedPrecisionArith(oper, arg1, arg2, context, target, checkOverflow context) in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordArith BuiltIns.ArithAdd, arg1, arg2=BICConstnt(value, _)}, context, _, destination, tailCode) = let val target = asTarget destination (* If the argument is a constant we can subtract the tag beforehand. N.B. it is possible to have type-incorrect values in dead code. i.e. code that will never be executed because of a run-time check. *) val constVal = if isShort value then semitag(Word.toLargeIntX(toShort value)) else 0 val (arg1Code, aReg1) = codeToPRegRev(arg1, context, tailCode) in (BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=aReg1, operand2=IntegerConstant constVal, ccRef = newCCRef(), opSize=polyWordOpSize}) :: arg1Code, RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordArith BuiltIns.ArithAdd, arg1=BICConstnt(value, _), arg2}, context, _, destination, tailCode) = let val target = asTarget destination (* If the argument is a constant we can subtract the tag beforehand. Check for short - see comment above. *) val constVal = if isShort value then semitag(Word.toLargeIntX(toShort value)) else 0 val (arg2Code, aReg2) = codeToPRegRev(arg2, context, tailCode) in (BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=aReg2, operand2=IntegerConstant constVal, ccRef = newCCRef(), opSize=polyWordOpSize}) :: arg2Code, RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordArith BuiltIns.ArithAdd, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) (* Use LEA to do the addition since we're not concerned with overflow. This is shorter than subtracting the tag and adding the values and also moves the result into the appropriate register. *) val code = arg1Code @ arg2Code @ [BlockSimple(LoadEffectiveAddress{base=SOME aReg1, offset= ~1, index=MemIndex1 aReg2, dest=target, opSize=polyWordOpSize})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordArith BuiltIns.ArithSub, arg1, arg2=BICConstnt(value, _)}, context, _, destination, tailCode) = let val target = asTarget destination (* If the argument is a constant we can subtract the tag beforehand. Check for short - see comment above. *) val constVal = if isShort value then semitag(Word.toLargeIntX(toShort value)) else 0 val (arg1Code, aReg1) = codeToPRegRev(arg1, context, tailCode) in (BlockSimple(ArithmeticFunction{oper=SUB, resultReg=target, operand1=aReg1, operand2=IntegerConstant constVal, ccRef=newCCRef(), opSize=polyWordOpSize}) :: arg1Code, RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordArith BuiltIns.ArithSub, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val aReg3 = newPReg() val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) val code = arg1Code @ arg2Code @ (* Do the subtraction and add in the tag bit. This could be reordered if we have cascaded operations since we don't need to check for overflow. *) [BlockSimple(ArithmeticFunction{oper=SUB, resultReg=aReg3, operand1=aReg1, operand2=RegisterArgument aReg2, ccRef=newCCRef(), opSize=polyWordOpSize}), BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=aReg3, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordArith BuiltIns.ArithMult, arg1, arg2=BICConstnt(value, _)}, context, _, destination, tailCode) = codeMultiplyConstantWordRev(arg1, context, destination, if isShort value then toShort value else 0w0, tailCode) | codeToICodeBinaryRev({oper=BuiltIns.WordArith BuiltIns.ArithMult, arg1=BICConstnt(value, _), arg2}, context, _, destination, tailCode) = codeMultiplyConstantWordRev(arg2, context, destination, if isShort value then toShort value else 0w0, tailCode) | codeToICodeBinaryRev({oper=BuiltIns.WordArith BuiltIns.ArithMult, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) val arg1Untagged = newUReg() and arg2Untagged = newUReg() and resUntagged = newUReg() val code = arg1Code @ arg2Code @ (* Shift one argument and subtract the tag from the other. It's possible this could be reordered if we have a value that is already untagged. *) [BlockSimple(UntagValue{source=aReg1, dest=arg1Untagged, isSigned=false, cache=NONE, opSize=polyWordOpSize}), BlockSimple(ArithmeticFunction{oper=SUB, resultReg=arg2Untagged, operand1=aReg2, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize}), BlockSimple(Multiplication{resultReg=resUntagged, operand1=arg1Untagged, operand2=RegisterArgument arg2Untagged, ccRef=newCCRef(), opSize=polyWordOpSize}), BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=resUntagged, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordArith BuiltIns.ArithDiv, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) val arg1Untagged = newUReg() and arg2Untagged = newUReg() val quotient = newUReg() and remainder = newUReg() val code = arg1Code @ arg2Code @ (* Shift both of the arguments to remove the tags. We don't test for zero here - that's done explicitly. *) [BlockSimple(UntagValue{source=aReg1, dest=arg1Untagged, isSigned=false, cache=NONE, opSize=polyWordOpSize}), BlockSimple(UntagValue{source=aReg2, dest=arg2Untagged, isSigned=false, cache=NONE, opSize=polyWordOpSize}), BlockSimple(Division { isSigned = false, dividend=arg1Untagged, divisor=RegisterArgument arg2Untagged, quotient=quotient, remainder=remainder, opSize=polyWordOpSize }), BlockSimple(TagValue { source=quotient, dest=target, isSigned=false, opSize=polyWordOpSize })] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordArith BuiltIns.ArithMod, arg1, arg2}, context, _, destination, tailCode) = let (* Identical to Quot except that the result is the remainder. *) val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) val arg1Untagged = newUReg() and arg2Untagged = newUReg() val quotient = newUReg() and remainder = newUReg() val code = arg1Code @ arg2Code @ (* Shift both of the arguments to remove the tags. *) [BlockSimple(UntagValue{source=aReg1, dest=arg1Untagged, isSigned=false, cache=NONE, opSize=polyWordOpSize}), BlockSimple(UntagValue{source=aReg2, dest=arg2Untagged, isSigned=false, cache=NONE, opSize=polyWordOpSize}), BlockSimple(Division { isSigned = false, dividend=arg1Untagged, divisor=RegisterArgument arg2Untagged, quotient=quotient, remainder=remainder, opSize=polyWordOpSize }), BlockSimple(TagValue { source=remainder, dest=target, isSigned=false, opSize=polyWordOpSize })] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordArith _, ...}, _, _, _, _) = raise InternalError "codeToICodeNonRev: WordArith - unimplemented operation" | codeToICodeBinaryRev({oper=BuiltIns.WordLogical logOp, arg1, arg2=BICConstnt(value, _)}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, arg1Reg) = codeToPReg(arg1, context) (* Use a semitagged value for XOR. This preserves the tag bit. Use toLargeIntX here because the operations will sign-extend 32-bit values. *) val constVal = if isShort value then (case logOp of BuiltIns.LogicalXor => semitag | _ => tag) (Word.toLargeIntX(toShort value)) else 0 val oper = case logOp of BuiltIns.LogicalOr => OR | BuiltIns.LogicalAnd => AND | BuiltIns.LogicalXor => XOR (* If we AND with a value that fits in 32-bits we can use a 32-bit operation. *) val opSize = if logOp = BuiltIns.LogicalAnd andalso constVal <= 0xffffffff andalso constVal >= 0 then OpSize32 else polyWordOpSize val code = arg1Code @ [BlockSimple(ArithmeticFunction{oper=oper, resultReg=target, operand1=arg1Reg, operand2=IntegerConstant constVal, ccRef=newCCRef(), opSize=opSize})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordLogical logOp, arg1=BICConstnt(value, _), arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg2Code, arg2Reg) = codeToPReg(arg2, context) (* Use a semitagged value for XOR. This preserves the tag bit. *) val constVal = if isShort value then (case logOp of BuiltIns.LogicalXor => semitag | _ => tag) (Word.toLargeIntX(toShort value)) else 0 val oper = case logOp of BuiltIns.LogicalOr => OR | BuiltIns.LogicalAnd => AND | BuiltIns.LogicalXor => XOR (* If we AND with a value that fits in 32-bits we can use a 32-bit operation. *) val opSize = if logOp = BuiltIns.LogicalAnd andalso constVal <= 0xffffffff andalso constVal >= 0 then OpSize32 else polyWordOpSize val code = arg2Code @ [BlockSimple(ArithmeticFunction{oper=oper, resultReg=target, operand1=arg2Reg, operand2=IntegerConstant constVal, ccRef=newCCRef(), opSize=opSize})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordLogical BuiltIns.LogicalOr, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, arg1Reg) = codeToPReg(arg1, context) val (arg2Code, arg2Reg) = codeToPReg(arg2, context) val code = arg1Code @ arg2Code @ (* Or-ing preserves the tag bit. *) [BlockSimple(ArithmeticFunction{oper=OR, resultReg=target, operand1=arg1Reg, operand2=RegisterArgument arg2Reg, ccRef=newCCRef(), opSize=polyWordOpSize})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordLogical BuiltIns.LogicalAnd, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, arg1Reg) = codeToPReg(arg1, context) val (arg2Code, arg2Reg) = codeToPReg(arg2, context) val code = arg1Code @ arg2Code @ (* Since they're both tagged the result will be tagged. *) [BlockSimple(ArithmeticFunction{oper=AND, resultReg=target, operand1=arg1Reg, operand2=RegisterArgument arg2Reg, ccRef=newCCRef(), opSize=polyWordOpSize})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordLogical BuiltIns.LogicalXor, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, arg1Reg) = codeToPReg(arg1, context) val (arg2Code, arg2Reg) = codeToPReg(arg2, context) val aReg3 = newPReg() val code = arg1Code @ arg2Code @ (* We need to restore the tag bit after the operation. *) [BlockSimple(ArithmeticFunction{oper=XOR, resultReg=aReg3, operand1=arg1Reg, operand2=RegisterArgument arg2Reg, ccRef=newCCRef(), opSize=polyWordOpSize}), BlockSimple(ArithmeticFunction{oper=OR, resultReg=target, operand1=aReg3, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.WordShift BuiltIns.ShiftLeft, arg1, arg2=BICConstnt(value, _)}, context, _, destination, tailCode) = (* Use the general case multiplication code. This will use a shift except for small values. It does detect special cases such as multiplication by 4 and 8 which can be implemented with LEA. *) codeMultiplyConstantWordRev(arg1, context, destination, if isShort value then Word.<<(0w1, toShort value) else 0w1, tailCode) | codeToICodeBinaryRev({oper=BuiltIns.WordShift shift, arg1, arg2}, context, _, destination, tailCode) = (* N.B. X86 shifts of greater than the word length mask the higher bits. That isn't what ML wants but that is dealt with at a higher level *) let open BuiltIns val target = asTarget destination (* Load the value into an untagged register. If this is a left shift we need to clear the tag bit. We don't need to do that for right shifts. *) val argRegUntagged = newUReg() val arg1Code = case arg1 of BICConstnt(value, _) => let (* Remove the tag bit. This isn't required for right shifts. *) val cnstntVal = if isShort value then semitag(Word.toLargeInt(toShort value)) else 1 in [BlockSimple(LoadArgument{source=IntegerConstant cnstntVal, dest=argRegUntagged, kind=movePolyWord})] end | _ => let val (arg1Code, arg1Reg) = codeToPReg(arg1, context) val removeTag = case shift of ShiftLeft => ArithmeticFunction{oper=SUB, resultReg=argRegUntagged, operand1=arg1Reg, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize} | _ => LoadArgument{source=RegisterArgument arg1Reg, dest=argRegUntagged, kind=movePolyWord} in arg1Code @ [BlockSimple removeTag] end (* The shift amount can usefully be a constant. *) val (arg2Code, untag2Code, arg2Arg) = codeAsUntaggedByte(arg2, false, context) val resRegUntagged = newUReg() val shiftOp = case shift of ShiftLeft => SHL | ShiftRightLogical => SHR | ShiftRightArithmetic => SAR val code = arg1Code @ arg2Code @ untag2Code @ [BlockSimple(ShiftOperation{ shift=shiftOp, resultReg=resRegUntagged, operand=argRegUntagged, shiftAmount=arg2Arg, ccRef=newCCRef(), opSize=polyWordOpSize }), (* Set the tag by ORing it in. This will work whether or not a right shift has shifted a 1 into this position. *) BlockSimple( ArithmeticFunction{oper=OR, resultReg=target, operand1=resRegUntagged, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.AllocateByteMemory, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val sizeReg = newPReg() and baseReg = newPReg() val sizeCode = codeToICodeTarget(arg1, context, false, sizeReg) val (flagsCode, flagUntag, flagArg) = codeAsUntaggedByte(arg2, false, context) val code =sizeCode @ flagsCode @ [BlockSimple(AllocateMemoryVariable{size=sizeReg, dest=baseReg, saveRegs=[]})] @ flagUntag @ [BlockSimple(StoreArgument{ source=flagArg, base=baseReg, offset= ~1, index=memIndexOrObject, kind=MoveByte, isMutable=false}), BlockSimple InitialisationComplete, BlockSimple(LoadArgument{ source=RegisterArgument baseReg, dest=target, kind=movePolyWord})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordComparison test, arg1, arg2}, context, _, destination, tailCode) = let val ccRef = newCCRef() val (arg1Code, arg1Reg) = codeToPRegRev(arg1, context, tailCode) (* In X64 we can extract the word from a constant and do the comparison directly. That can't be done in X86/32 because the value isn't tagged and might look like an address. The RTS scans for comparisons with inline constant addresses. *) val (arg2Code, arg2Operand) = if targetArch <> Native32Bit then (* Native 64-bit or 32-in-64. *) ( case arg2 of BICConstnt(value, _) => (arg1Code, IntegerConstant(largeWordConstant value)) | _ => let val (code, reg) = codeToPRegRev(arg2, context, arg1Code) in (code, wordAt reg) end ) else let val (code, reg) = codeToPRegRev(arg2, context, arg1Code) in (code, wordAt reg) end val argReg = newUReg() val target = asTarget destination val code = makeBoolResultRev(testAsBranch(test, false, true), ccRef, target, BlockSimple(WordComparison{arg1=argReg, arg2=arg2Operand, ccRef=ccRef, opSize=nativeWordOpSize}) :: BlockSimple(LoadArgument{source=wordAt arg1Reg, dest=argReg, kind=moveNativeWord}) :: arg2Code) in (code, RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordArith BuiltIns.ArithAdd, arg1, arg2=BICConstnt(value, _)}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val aReg3 = newUReg() val argReg = newUReg() val constantValue = largeWordConstant value val code =arg1Code @ [BlockSimple(LoadArgument{source=wordAt aReg1, dest=argReg, kind=moveNativeWord}), BlockSimple(ArithmeticFunction{oper=ADD, resultReg=aReg3, operand1=argReg, operand2=IntegerConstant constantValue, ccRef=newCCRef(), opSize=nativeWordOpSize}), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=aReg3, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordArith BuiltIns.ArithAdd, arg1=BICConstnt(value, _), arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg2Code, aReg2) = codeToPReg(arg2, context) val aReg3 = newUReg() val argReg = newUReg() val constantValue = largeWordConstant value val code = arg2Code @ [BlockSimple(LoadArgument{source=wordAt aReg2, dest=argReg, kind=moveNativeWord}), BlockSimple(ArithmeticFunction{oper=ADD, resultReg=aReg3, operand1=argReg, operand2=IntegerConstant constantValue, ccRef=newCCRef(), opSize=nativeWordOpSize}), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=aReg3, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordArith BuiltIns.ArithAdd, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) val aReg3 = newUReg() val argReg = newUReg() val code = arg1Code @ arg2Code @ [BlockSimple(LoadArgument{source=wordAt aReg1, dest=argReg, kind=moveNativeWord}), BlockSimple(ArithmeticFunction{oper=ADD, resultReg=aReg3, operand1=argReg, operand2=wordAt aReg2, ccRef=newCCRef(), opSize=nativeWordOpSize}), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=aReg3, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordArith BuiltIns.ArithSub, arg1, arg2=BICConstnt(value, _)}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val aReg3 = newUReg() val argReg = newUReg() val constantValue = largeWordConstant value val code = arg1Code @ [BlockSimple(LoadArgument{source=wordAt aReg1, dest=argReg, kind=moveNativeWord}), BlockSimple(ArithmeticFunction{oper=SUB, resultReg=aReg3, operand1=argReg, operand2=IntegerConstant constantValue, ccRef=newCCRef(), opSize=nativeWordOpSize}), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=aReg3, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordArith BuiltIns.ArithSub, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) val aReg3 = newUReg() val argReg = newUReg() val code = arg1Code @ arg2Code @ [BlockSimple(LoadArgument{source=wordAt aReg1, dest=argReg, kind=moveNativeWord}), BlockSimple(ArithmeticFunction{oper=SUB, resultReg=aReg3, operand1=argReg, operand2=wordAt aReg2, ccRef=newCCRef(), opSize=nativeWordOpSize}), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=aReg3, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordArith BuiltIns.ArithMult, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val resValue = newUReg() val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) val argReg1 = newUReg() val code = arg1Code @ arg2Code @ [BlockSimple(LoadArgument{source=wordAt aReg1, dest=argReg1, kind=moveNativeWord}), BlockSimple(Multiplication{resultReg=resValue, operand1=argReg1, operand2=wordAt aReg2, ccRef=newCCRef(), opSize=nativeWordOpSize}), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=resValue, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordArith BuiltIns.ArithDiv, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) val quotient = newUReg() and remainder = newUReg() val dividendReg = newUReg() and divisorReg = newUReg() val code = arg1Code @ arg2Code @ (* We don't test for zero here - that's done explicitly. *) [BlockSimple(LoadArgument{source=wordAt aReg1, dest=dividendReg, kind=moveNativeWord}), BlockSimple(LoadArgument{source=wordAt aReg2, dest=divisorReg, kind=moveNativeWord}), BlockSimple(Division { isSigned = false, dividend=dividendReg, divisor=RegisterArgument divisorReg, quotient=quotient, remainder=remainder, opSize=nativeWordOpSize }), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=quotient, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordArith BuiltIns.ArithMod, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) val quotient = newUReg() and remainder = newUReg() val dividendReg = newUReg() and divisorReg = newUReg() val code = arg1Code @ arg2Code @ (* We don't test for zero here - that's done explicitly. *) [BlockSimple(LoadArgument{source=wordAt aReg1, dest=dividendReg, kind=moveNativeWord}), BlockSimple(LoadArgument{source=wordAt aReg2, dest=divisorReg, kind=moveNativeWord}), BlockSimple(Division { isSigned = false, dividend=dividendReg, divisor=RegisterArgument divisorReg, quotient=quotient, remainder=remainder, opSize=nativeWordOpSize }), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=remainder, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordArith _, ...}, _, _, _, _) = raise InternalError "codeToICodeNonRev: LargeWordArith - unimplemented operation" | codeToICodeBinaryRev({oper=BuiltIns.LargeWordLogical logOp, arg1, arg2=BICConstnt(value, _)}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val aReg3 = newUReg() val argReg = newUReg() val constantValue = largeWordConstant value val oper = case logOp of BuiltIns.LogicalOr => OR | BuiltIns.LogicalAnd => AND | BuiltIns.LogicalXor => XOR (* If we AND with a value that fits in 32-bits we can use a 32-bit operation. *) val opSize = if logOp = BuiltIns.LogicalAnd andalso constantValue <= 0xffffffff andalso constantValue >= 0 then OpSize32 else nativeWordOpSize val code = arg1Code @ [BlockSimple(LoadArgument{source=wordAt aReg1, dest=argReg, kind=moveNativeWord}), BlockSimple(ArithmeticFunction{oper=oper, resultReg=aReg3, operand1=argReg, operand2=IntegerConstant constantValue, ccRef=newCCRef(), opSize=opSize}), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=aReg3, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordLogical logOp, arg1=BICConstnt(value, _), arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg2Code, aReg2) = codeToPReg(arg2, context) val aReg3 = newUReg() val argReg = newUReg() val constantValue = largeWordConstant value val oper = case logOp of BuiltIns.LogicalOr => OR | BuiltIns.LogicalAnd => AND | BuiltIns.LogicalXor => XOR (* If we AND with a value that fits in 32-bits we can use a 32-bit operation. *) val opSize = if logOp = BuiltIns.LogicalAnd andalso constantValue <= 0xffffffff andalso constantValue >= 0 then OpSize32 else nativeWordOpSize val code = arg2Code @ [BlockSimple(LoadArgument{source=wordAt aReg2, dest=argReg, kind=moveNativeWord}), BlockSimple(ArithmeticFunction{oper=oper, resultReg=aReg3, operand1=argReg, operand2=IntegerConstant constantValue, ccRef=newCCRef(), opSize=opSize}), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=aReg3, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordLogical logOp, arg1, arg2}, context, _, destination, tailCode) = let val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) val aReg3 = newUReg() val argReg = newUReg() val oper = case logOp of BuiltIns.LogicalOr => OR | BuiltIns.LogicalAnd => AND | BuiltIns.LogicalXor => XOR val code = arg1Code @ arg2Code @ [BlockSimple(LoadArgument{source=wordAt aReg1, dest=argReg, kind=moveNativeWord}), BlockSimple(ArithmeticFunction{oper=oper, resultReg=aReg3, operand1=argReg, operand2=wordAt aReg2, ccRef=newCCRef(), opSize=nativeWordOpSize}), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=aReg3, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.LargeWordShift shift, arg1, arg2}, context, _, destination, tailCode) = (* The shift is always a Word.word value i.e. tagged. There is a check at the higher level that the shift does not exceed 32/64 bits. *) let open BuiltIns val target = asTarget destination val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, untag2Code, arg2Arg) = codeAsUntaggedByte(arg2, false, context) val aReg3 = newUReg() val shiftOp = case shift of ShiftLeft => SHL | ShiftRightLogical => SHR | ShiftRightArithmetic => SAR val argReg = newUReg() val code = arg1Code @ arg2Code @ [BlockSimple(LoadArgument{source=wordAt aReg1, dest=argReg, kind=moveNativeWord})] @ untag2Code @ [BlockSimple(ShiftOperation{ shift=shiftOp, resultReg=aReg3, operand=argReg, shiftAmount=arg2Arg, ccRef=newCCRef(), opSize=nativeWordOpSize }), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=aReg3, dest=target, saveRegs=[]})] in (revApp(code, tailCode), RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.RealArith(fpOpPrec as (fpOp, fpPrec)), arg1, arg2}, context, _, destination, tailCode) = let open BuiltIns val commutative = case fpOp of ArithSub => NonCommutative | ArithDiv => NonCommutative | ArithAdd => Commutative | ArithMult => Commutative | _ => raise InternalError "codeToICodeNonRev: RealArith - unimplemented operation" val (argCodeRev, fpRegSrc, arg2Value) = codeFPBinaryArgsRev(arg1, arg2, fpPrec, commutative, context, []) val argCode = List.rev argCodeRev val target = asTarget destination val fpRegDest = newUReg() val arith = case fpMode of FPModeX87 => let val fpOp = case fpOp of ArithAdd => FADD | ArithSub => FSUB | ArithMult => FMUL | ArithDiv => FDIV | _ => raise InternalError "codeToICodeNonRev: RealArith - unimplemented operation" val isDouble = case fpPrec of PrecSingle => false | PrecDouble => true in [BlockSimple(X87FPArith{ opc=fpOp, resultReg=fpRegDest, arg1=fpRegSrc, arg2=arg2Value, isDouble=isDouble})] end | FPModeSSE2 => let val fpOp = case fpOpPrec of (ArithAdd, PrecSingle) => SSE2BAddSingle | (ArithSub, PrecSingle) => SSE2BSubSingle | (ArithMult, PrecSingle) => SSE2BMulSingle | (ArithDiv, PrecSingle) => SSE2BDivSingle | (ArithAdd, PrecDouble) => SSE2BAddDouble | (ArithSub, PrecDouble) => SSE2BSubDouble | (ArithMult, PrecDouble) => SSE2BMulDouble | (ArithDiv, PrecDouble) => SSE2BDivDouble | _ => raise InternalError "codeToICodeNonRev: RealArith - unimplemented operation" in [BlockSimple(SSE2FPBinary{ opc=fpOp, resultReg=fpRegDest, arg1=fpRegSrc, arg2=arg2Value})] end (* Box or tag the result. *) val result = boxOrTagReal(fpRegDest, target, fpPrec) in (revApp(argCode @ arith @ result, tailCode), RegisterArgument target, false) end (* Floating point comparison. This is complicated because we have different instruction sequences for SSE2 and X87. We also have to get the handling of unordered (NaN) values right. All the tests are treated as false if either argument is a NaN. To combine that test with the other tests we sometimes have to reverse the comparison. *) | codeToICodeBinaryRev({oper=BuiltIns.RealComparison(BuiltIns.TestEqual, precision), arg1, arg2}, context, _, destination, tailCode) = let (* Get the arguments. It's commutative. *) val (arg2Code, fpReg, arg2Val) = codeFPBinaryArgsRev(arg1, arg2, precision, Commutative, context, tailCode) val ccRef1 = newCCRef() and ccRef2 = newCCRef() val testReg1 = newUReg() and testReg2 = newUReg() and testReg3 = newUReg() (* If this is X87 we get the condition into RAX and test it there. If it is SSE2 we have to treat the unordered result (parity set) specially. *) val isDouble = precision = BuiltIns.PrecDouble val target = asTarget destination val code = case fpMode of FPModeX87 => makeBoolResultRev(JE, ccRef2, target, BlockSimple(ArithmeticFunction{ oper=XOR, resultReg=testReg3, operand1=testReg2, operand2=IntegerConstant 0x4000, ccRef=ccRef2, opSize=OpSize32 }) :: BlockSimple(ArithmeticFunction{ oper=AND, resultReg=testReg2, operand1=testReg1, operand2=IntegerConstant 0x4400, ccRef=newCCRef(), opSize=OpSize32 }) :: BlockSimple(X87FPGetCondition { ccRef=ccRef1, dest=testReg1 }) :: BlockSimple(X87Compare{arg1=fpReg, arg2=arg2Val, ccRef=ccRef1, isDouble = isDouble}) :: arg2Code) | FPModeSSE2 => let val noParityLabel = newLabel() val resultLabel = newLabel() val falseLabel = newLabel() val trueLabel = newLabel() val mergeReg = newMergeReg() in BlockSimple(LoadArgument{ source=RegisterArgument mergeReg, dest=target, kind=Move32Bit }) :: BlockLabel resultLabel :: BlockFlow(Unconditional resultLabel) :: (* Result is false if parity is set i.e. unordered or if unequal. *) BlockSimple(LoadArgument{ source=IntegerConstant(tag 0), dest=mergeReg, kind=Move32Bit }) :: BlockLabel falseLabel :: BlockFlow(Unconditional resultLabel) :: (* Result is true if it's ordered and equal. *) BlockSimple(LoadArgument{ source=IntegerConstant(tag 1), dest=mergeReg, kind=Move32Bit }) :: BlockLabel trueLabel :: (* Not unordered - test the equality *) BlockFlow(Conditional{ccRef=ccRef1, condition=JE, trueJump=trueLabel, falseJump=falseLabel}) :: BlockLabel noParityLabel :: (* Go to falseLabel if unordered and therefore not equal. *) BlockFlow(Conditional{ccRef=ccRef1, condition=JP, trueJump=falseLabel, falseJump=noParityLabel}) :: BlockSimple(SSE2Compare{arg1=fpReg, arg2=arg2Val, ccRef=ccRef1, isDouble = isDouble}) :: arg2Code end in (code, RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.RealComparison(BuiltIns.TestUnordered, precision), arg1, arg2}, context, _, destination, tailCode) = let (* The unordered test is really included because it is easy to implement and is the simplest way of implementing isNan. *) (* Get the arguments. It's commutative. *) val (arg2Code, fpReg, arg2Val) = codeFPBinaryArgsRev(arg1, arg2, precision, Commutative, context, tailCode) val ccRef1 = newCCRef() and ccRef2 = newCCRef() val testReg1 = newUReg() and testReg2 = newUReg() and testReg3 = newUReg() (* If this is X87 we get the condition into RAX and test it there. If it is SSE2 we have to treat the unordered result (parity set) specially. *) val isDouble = precision = BuiltIns.PrecDouble val target = asTarget destination val code = case fpMode of FPModeX87 => (* And with 0x4500. We have to use XOR rather than CMP to avoid having an untagged constant comparison. *) makeBoolResultRev(JE, ccRef2, target, BlockSimple(ArithmeticFunction{ oper=XOR, resultReg=testReg3, operand1=testReg2, operand2=IntegerConstant 0x4500, ccRef=ccRef2, opSize=OpSize32 }) :: BlockSimple(ArithmeticFunction{ oper=AND, resultReg=testReg2, operand1=testReg1, operand2=IntegerConstant 0x4500, ccRef=newCCRef(), opSize=OpSize32 }) :: BlockSimple(X87FPGetCondition { ccRef=ccRef1, dest=testReg1 }) :: BlockSimple(X87Compare{arg1=fpReg, arg2=arg2Val, ccRef=ccRef1, isDouble = isDouble}) :: arg2Code) | FPModeSSE2 => makeBoolResultRev(JP, ccRef1, target, BlockSimple(SSE2Compare{arg1=fpReg, arg2=arg2Val, ccRef=ccRef1, isDouble = isDouble}) :: arg2Code) in (code, RegisterArgument target, false) end | codeToICodeBinaryRev({oper=BuiltIns.RealComparison(comparison, precision), arg1, arg2}, context, _, destination, tailCode) = let (* Ordered comparisons are complicated because they are all defined to be false if either argument is a NaN. We have two different tests for a > b and a >= b and implement a < b and a <= b by changing the order of the arguments. *) val (arg1Code, arg1Value) = codeFPArgument(arg1, precision, context, tailCode) val (arg2Code, arg2Value) = codeFPArgument(arg2, precision, context, arg1Code) val (regArg, opArg, isGeq) = case comparison of BuiltIns.TestGreater => (arg1Value, arg2Value, false) | BuiltIns.TestLess => (arg2Value, arg1Value, false) (* Reversed: aa. *) | BuiltIns.TestGreaterEqual => (arg1Value, arg2Value, true) | BuiltIns.TestLessEqual => (arg2Value, arg1Value, true) (* Reversed: a<=b is b>=a. *) | _ => raise InternalError "RealComparison: unimplemented operation" (* Load the first operand into a register. *) val (fpReg, loadCode) = case regArg of RegisterArgument fpReg => (fpReg, arg2Code) | regArg => let val fpReg = newUReg() val moveOp = case precision of BuiltIns.PrecDouble => MoveDouble | BuiltIns.PrecSingle => MoveFloat in (fpReg, BlockSimple(LoadArgument{source=regArg, dest=fpReg, kind=moveOp}) :: arg2Code) end val isDouble = precision = BuiltIns.PrecDouble val target = asTarget destination val code = case fpMode of FPModeX87 => let val testReg1 = newUReg() and testReg2 = newUReg() val ccRef1 = newCCRef() and ccRef2 = newCCRef() val testBits = if isGeq then 0x500 else 0x4500 in makeBoolResultRev(JE, ccRef2, target, BlockSimple(ArithmeticFunction{ oper=AND, resultReg=testReg2, operand1=testReg1, operand2=IntegerConstant testBits, ccRef=ccRef2, opSize=OpSize32 }) :: BlockSimple(X87FPGetCondition { ccRef=ccRef1, dest=testReg1 }) :: BlockSimple(X87Compare{arg1=fpReg, arg2=opArg, ccRef=ccRef1, isDouble = isDouble}) :: loadCode) end | FPModeSSE2 => let val ccRef1 = newCCRef() val condition = if isGeq then JNB (* >=, <= *) else JA (* >, < *) in makeBoolResultRev(condition, ccRef1, target, BlockSimple(SSE2Compare{arg1=fpReg, arg2=opArg, ccRef=ccRef1, isDouble = isDouble}) :: loadCode) end in (code, RegisterArgument target, false) end (* Multiply tagged word by a constant. We're not concerned with overflow so it's possible to use various short cuts. *) and codeMultiplyConstantWordRev(arg, context, destination, multiplier, tailCode) = let val target = asTarget destination val (argCode, aReg) = codeToPReg(arg, context) val doMultiply = case multiplier of 0w0 => [BlockSimple(LoadArgument{source=IntegerConstant 1, dest=target, kind=movePolyWord})] | 0w1 => [BlockSimple(LoadArgument{source=RegisterArgument aReg, dest=target, kind=movePolyWord})] | 0w2 => [BlockSimple(LoadEffectiveAddress{base=SOME aReg, offset= ~1, index=MemIndex1 aReg, dest=target, opSize=polyWordOpSize})] | 0w3 => [BlockSimple(LoadEffectiveAddress{base=SOME aReg, offset= ~2, index=MemIndex2 aReg, dest=target, opSize=polyWordOpSize})] | 0w4 => [BlockSimple(LoadEffectiveAddress{base=NONE, offset= ~3, index=MemIndex4 aReg, dest=target, opSize=polyWordOpSize})] | 0w5 => [BlockSimple(LoadEffectiveAddress{base=SOME aReg, offset= ~4, index=MemIndex4 aReg, dest=target, opSize=polyWordOpSize})] | 0w8 => [BlockSimple(LoadEffectiveAddress{base=NONE, offset= ~7, index=MemIndex8 aReg, dest=target, opSize=polyWordOpSize})] | 0w9 => [BlockSimple(LoadEffectiveAddress{base=SOME aReg, offset= ~8, index=MemIndex8 aReg, dest=target, opSize=polyWordOpSize})] | _ => let val tReg = newUReg() val tagCorrection = Word.toLargeInt multiplier - 1 fun getPower2 n = let fun p2 (n, l) = if n = 0w1 then SOME l else if Word.andb(n, 0w1) = 0w1 then NONE else p2(Word.>>(n, 0w1), l+0w1) in if n = 0w0 then NONE else p2(n,0w0) end val multiply = case getPower2 multiplier of SOME power => (* Shift it including the tag. *) BlockSimple(ShiftOperation{ shift=SHL, resultReg=tReg, operand=aReg, shiftAmount=IntegerConstant(Word.toLargeInt power), ccRef=newCCRef(), opSize=polyWordOpSize }) | NONE => (* Multiply including the tag. *) BlockSimple(Multiplication{resultReg=tReg, operand1=aReg, operand2=IntegerConstant(Word.toLargeInt multiplier), ccRef=newCCRef(), opSize=polyWordOpSize}) in [multiply, BlockSimple(ArithmeticFunction{oper=SUB, resultReg=target, operand1=tReg, operand2=IntegerConstant tagCorrection, ccRef=newCCRef(), opSize=polyWordOpSize})] end in (revApp(argCode @ doMultiply, tailCode), RegisterArgument target, false) end and codeToICodeAllocate({numWords as BICConstnt(length, _), flags as BICConstnt(flagValue, _), initial}, context, _, destination) = (* Constant length and flags is used for ref. We could handle other cases. *) if isShort length andalso isShort flagValue andalso toShort length = 0w1 then let val target = asTarget destination (* Force a different register. *) val vecLength = Word.toInt(toShort length) val flagByte = Word8.fromLargeWord(Word.toLargeWord(toShort flagValue)) val memAddr = newPReg() and valueReg = newPReg() fun initialise n = BlockSimple(StoreArgument{ source=RegisterArgument valueReg, offset=n*Word.toInt wordSize, base=memAddr, index=memIndexOrObject, kind=movePolyWord, isMutable=false}) val code = codeToICodeTarget(initial, context, false, valueReg) @ [BlockSimple(AllocateMemoryOperation{size=vecLength, flags=flagByte, dest=memAddr, saveRegs=[]})] @ List.tabulate(vecLength, initialise) @ [BlockSimple InitialisationComplete, BlockSimple(LoadArgument{source=RegisterArgument memAddr, dest=target, kind=movePolyWord})] in (code, RegisterArgument target, false) end else (* If it's longer use the full run-time form. *) allocateMemoryVariable(numWords, flags, initial, context, destination) | codeToICodeAllocate({numWords, flags, initial}, context, _, destination) = allocateMemoryVariable(numWords, flags, initial, context, destination) and codeToICodeLoad({kind=LoadStoreMLWord _, address}, context, _, destination) = let val target = asTarget destination val (codeBaseIndex, codeUntag, memLoc) = codeAddress(address, false, context) in (codeBaseIndex @ codeUntag @ [BlockSimple(LoadArgument {source=MemoryLocation memLoc, dest=target, kind=movePolyWord})], RegisterArgument target, false) end | codeToICodeLoad({kind=LoadStoreMLByte _, address}, context, _, destination) = let val target = asTarget destination val (codeBaseIndex, codeUntag, memLoc) = codeAddress(address, true, context) val untaggedResReg = newUReg() in (codeBaseIndex @ codeUntag @ [BlockSimple(LoadArgument { source=MemoryLocation memLoc, dest=untaggedResReg, kind=MoveByte}), BlockSimple(TagValue {source=untaggedResReg, dest=target, isSigned=false, opSize=OpSize32})], RegisterArgument target, false) end | codeToICodeLoad({kind=LoadStoreC8, address}, context, _, destination) = let (* Load a byte from C memory. This is almost exactly the same as LoadStoreMLByte except that the base address is a LargeWord.word value. *) val target = asTarget destination val (codeBaseIndex, codeUntag, memLoc) = codeCAddress(address, 0w1, context) val untaggedResReg = newUReg() in (codeBaseIndex @ codeUntag @ [BlockSimple(LoadArgument { source=MemoryLocation memLoc, dest=untaggedResReg, kind=MoveByte}), BlockSimple(TagValue {source=untaggedResReg, dest=target, isSigned=false, opSize=OpSize32})], RegisterArgument target, false) end | codeToICodeLoad({kind=LoadStoreC16, address}, context, _, destination) = let (* Load a 16-bit value from C memory. *) val target = asTarget destination val (codeBaseIndex, codeUntag, memLoc) = codeCAddress(address, 0w2, context) val untaggedResReg = newUReg() in (codeBaseIndex @ codeUntag @ [BlockSimple(LoadArgument { source=MemoryLocation memLoc, dest=untaggedResReg, kind=Move16Bit}), BlockSimple(TagValue {source=untaggedResReg, dest=target, isSigned=false, opSize=OpSize32})], RegisterArgument target, false) end | codeToICodeLoad({kind=LoadStoreC32, address}, context, _, destination) = let (* Load a 32-bit value from C memory. If this is 64-bit mode we can tag it but if this is 32-bit mode we need to box it. *) val target = asTarget destination val (codeBaseIndex, codeUntag, memLoc) = codeCAddress(address, 0w4, context) val untaggedResReg = newUReg() val boxTagCode = if targetArch = Native64Bit then BlockSimple(TagValue {source=untaggedResReg, dest=target, isSigned=false, opSize=OpSize64 (* It becomes 33 bits *)}) else BlockSimple(BoxValue{boxKind=BoxLargeWord, source=untaggedResReg, dest=target, saveRegs=[]}) in (codeBaseIndex @ codeUntag @ [BlockSimple(LoadArgument { source=MemoryLocation memLoc, dest=untaggedResReg, kind=Move32Bit}), boxTagCode], RegisterArgument target, false) end | codeToICodeLoad({kind=LoadStoreC64, address}, context, _, destination) = let (* Load a 64-bit value from C memory. This is only allowed in 64-bit mode. The result is a boxed value. *) val _ = targetArch <> Native32Bit orelse raise InternalError "codeToICodeNonRev: BICLoadOperation LoadStoreC64 in 32-bit" val target = asTarget destination val (codeBaseIndex, codeUntag, memLoc) = codeCAddress(address, 0w8, context) val untaggedResReg = newUReg() in (codeBaseIndex @ codeUntag @ [BlockSimple(LoadArgument { source=MemoryLocation memLoc, dest=untaggedResReg, kind=Move64Bit}), BlockSimple(BoxValue{boxKind=BoxLargeWord, source=untaggedResReg, dest=target, saveRegs=[]})], RegisterArgument target, false) end | codeToICodeLoad({kind=LoadStoreCFloat, address}, context, _, destination) = let val target = asTarget destination val (codeBaseIndex, codeUntag, memLoc) = codeCAddress(address, 0w4, context) val untaggedResReg = newUReg() val boxFloat = case fpMode of FPModeX87 => BoxX87Double | FPModeSSE2 => BoxSSE2Double (* We need to convert the float into a double. *) val loadArg = case fpMode of FPModeX87 => BlockSimple(LoadArgument { source=MemoryLocation memLoc, dest=untaggedResReg, kind=MoveFloat}) | FPModeSSE2 => BlockSimple(SSE2FPUnary { source=MemoryLocation memLoc, resultReg=untaggedResReg, opc=SSE2UFloatToDouble}) in (codeBaseIndex @ codeUntag @ [loadArg, BlockSimple(BoxValue{boxKind=boxFloat, source=untaggedResReg, dest=target, saveRegs=[]})], RegisterArgument target, false) end | codeToICodeLoad({kind=LoadStoreCDouble, address}, context, _, destination) = let val target = asTarget destination val (codeBaseIndex, codeUntag, memLoc) = codeCAddress(address, 0w8, context) val untaggedResReg = newUReg() val boxFloat = case fpMode of FPModeX87 => BoxX87Double | FPModeSSE2 => BoxSSE2Double in (codeBaseIndex @ codeUntag @ [BlockSimple(LoadArgument { source=MemoryLocation memLoc, dest=untaggedResReg, kind=MoveDouble}), BlockSimple(BoxValue{boxKind=boxFloat, source=untaggedResReg, dest=target, saveRegs=[]})], RegisterArgument target, false) end | codeToICodeLoad({kind=LoadStoreUntaggedUnsigned, address}, context, _, destination) = let val target = asTarget destination val (codeBaseIndex, codeUntag, memLoc) = codeAddress(address, false, context) val untaggedResReg = newUReg() in (codeBaseIndex @ codeUntag @ [BlockSimple(LoadArgument { source=MemoryLocation memLoc, dest=untaggedResReg, kind=movePolyWord}), BlockSimple(TagValue {source=untaggedResReg, dest=target, isSigned=false, opSize=polyWordOpSize})], RegisterArgument target, false) end and codeToICodeStore({kind=LoadStoreMLWord _, address, value}, context, _, destination) = let val (sourceCode, source, _) = codeToICode(value, context, false, Allowed allowInMemMove) val (codeBaseIndex, codeUntag, {base, offset, index, ...}) = codeAddress(address, false, context) val code = codeBaseIndex @ sourceCode @ codeUntag @ [BlockSimple(StoreArgument {source=source, base=base, offset=offset, index=index, kind=movePolyWord, isMutable=true})] in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeStore({kind=LoadStoreMLByte _, address, value}, context, _, destination) = let val (codeBaseIndex, codeUntag, {base, offset, index, ...}) = codeAddress(address, true, context) (* We have to untag the value to store. *) val (valueCode, untagValue, valueArg) = codeAsUntaggedByte(value, false, context) val code = codeBaseIndex @ valueCode @ untagValue @ codeUntag @ [BlockSimple(StoreArgument {source=valueArg, base=base, offset=offset, index=index, kind=MoveByte, isMutable=true})] in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeStore({kind=LoadStoreC8, address, value}, context, _, destination) = let (* Store a byte to C memory. Almost exactly the same as LoadStoreMLByte. *) val (codeBaseIndex, codeUntag, {base, offset, index, ...}) = codeCAddress(address, 0w1, context) val (valueCode, untagValue, valueArg) = codeAsUntaggedByte(value, false, context) val code = codeBaseIndex @ valueCode @ untagValue @ codeUntag @ [BlockSimple(StoreArgument {source=valueArg, base=base, offset=offset, index=index, kind=MoveByte, isMutable=true})] in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeStore({kind=LoadStoreC16, address, value}, context, _, destination) = let (* Store a 16-bit value to C memory. *) val (codeBaseIndex, codeUntag, {base, offset, index, ...}) = codeCAddress(address, 0w2, context) (* We don't currently implement 16-bit constant moves so this must always be in a reg. *) val (valueCode, untagValue, valueArg) = codeAsUntaggedToReg(value, false, context) val code = codeBaseIndex @ valueCode @ untagValue @ codeUntag @ [BlockSimple(StoreArgument {source=RegisterArgument valueArg, base=base, offset=offset, index=index, kind=Move16Bit, isMutable=true})] in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeStore({kind=LoadStoreC32, address, value}, context, _, destination) = (* Store a 32-bit value. If this is 64-bit mode we untag it but if this is 32-bit mode we unbox it. *) let val (codeBaseIndex, codeUntag, {base, offset, index, ...}) = codeCAddress(address, 0w4, context) val code = if targetArch = Native64Bit then let (* We don't currently implement 32-bit constant moves so this must always be in a reg. *) val (valueCode, untagValue, valueArg) = codeAsUntaggedToReg(value, false, context) in codeBaseIndex @ valueCode @ untagValue @ codeUntag @ [BlockSimple(StoreArgument {source=RegisterArgument valueArg, base=base, offset=offset, index=index, kind=Move32Bit, isMutable=true})] end else let val (valueCode, valueReg) = codeToPReg(value, context) val valueReg1 = newUReg() in codeBaseIndex @ valueCode @ BlockSimple(LoadArgument{source=wordAt valueReg, dest=valueReg1, kind=Move32Bit}) :: codeUntag @ [BlockSimple(StoreArgument {source=RegisterArgument valueReg1, base=base, offset=offset, index=index, kind=Move32Bit, isMutable=true})] end in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeStore({kind=LoadStoreC64, address, value}, context, _, destination) = let (* Store a 64-bit value. *) val _ = targetArch <> Native32Bit orelse raise InternalError "codeToICodeNonRev: BICStoreOperation LoadStoreC64 in 32-bit" val (valueCode, valueReg) = codeToPReg(value, context) val valueReg1 = newUReg() val (codeBaseIndex, codeUntag, {base, offset, index, ...}) = codeCAddress(address, 0w8, context) val code = codeBaseIndex @ valueCode @ codeUntag @ [BlockSimple(LoadArgument{source=wordAt valueReg, dest=valueReg1, kind=Move64Bit}), BlockSimple(StoreArgument {source=RegisterArgument valueReg1, base=base, offset=offset, index=index, kind=Move64Bit, isMutable=true})] in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeStore({kind=LoadStoreCFloat, address, value}, context, _, destination) = let val floatReg = newUReg() and float2Reg = newUReg() val (codeBaseIndex, codeUntag, {base, offset, index, ...}) = codeCAddress(address, 0w4, context) val (valueCode, valueReg) = codeToPReg(value, context) (* If we're using an SSE2 reg we have to convert it from double to single precision. *) val (storeReg, cvtCode) = case fpMode of FPModeSSE2 => (float2Reg, [BlockSimple(SSE2FPUnary{opc=SSE2UDoubleToFloat, resultReg=float2Reg, source=RegisterArgument floatReg})]) | FPModeX87 => (floatReg, []) val code = codeBaseIndex @ valueCode @ codeUntag @ BlockSimple(LoadArgument{source=wordAt valueReg, dest=floatReg, kind=MoveDouble}) :: cvtCode @ [BlockSimple(StoreArgument {source=RegisterArgument storeReg, base=base, offset=offset, index=index, kind=MoveFloat, isMutable=true})] in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeStore({kind=LoadStoreCDouble, address, value}, context, _, destination) = let val floatReg = newUReg() val (codeBaseIndex, codeUntag, {base, offset, index, ...}) = codeCAddress(address, 0w8, context) val (valueCode, valueReg) = codeToPReg(value, context) val code = codeBaseIndex @ valueCode @ codeUntag @ [BlockSimple(LoadArgument{source=wordAt valueReg, dest=floatReg, kind=MoveDouble}), BlockSimple(StoreArgument {source=RegisterArgument floatReg, base=base, offset=offset, index=index, kind=MoveDouble, isMutable=true})] in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | codeToICodeStore({kind=LoadStoreUntaggedUnsigned, address, value}, context, _, destination) = let (* We have to untag the value to store. *) val (codeBaseIndex, codeUntag, {base, offset, index, ...}) = codeAddress(address, false, context) (* See if it's a constant. This is frequently used to set the last word of a string to zero. *) (* We have to be a bit more careful on the X86. We use moves to store constants that can include addresses. To avoid problems we only use a move if the value is zero or odd and so looks like a tagged value. *) val storeAble = case value of BICConstnt(value, _) => if not(isShort value) then NONE else let val ival = Word.toLargeIntX(toShort value) in if targetArch = Native64Bit then if is32bit ival then SOME ival else NONE else if ival = 0 orelse ival mod 2 = 1 then SOME ival else NONE end | _ => NONE val (storeVal, valCode) = case storeAble of SOME value => (IntegerConstant value (* Leave untagged *), []) | NONE => let val valueReg = newPReg() and valueReg1 = newUReg() in (RegisterArgument valueReg1, codeToICodeTarget(value, context, false, valueReg) @ [BlockSimple(UntagValue{dest=valueReg1, source=valueReg, isSigned=false, cache=NONE, opSize=polyWordOpSize})]) end val code = codeBaseIndex @ valCode @ codeUntag @ [BlockSimple(StoreArgument {source=storeVal, base=base, offset=offset, index=index, kind=movePolyWord, isMutable=true})] in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end and codeToICodeBlock({kind=BlockOpCompareByte, sourceLeft, destRight, length}, context, _, destination) = let (* This is effectively a big-endian comparison since we compare the bytes until we find an inequality. *) val target = asTarget destination val mergeResult = newMergeReg() val vec1Reg = newUReg() and vec2Reg = newUReg() val (leftCode, leftUntag, {base=leftBase, offset=leftOffset, index=leftIndex, ...}) = codeAddress(sourceLeft, true, context) val (rightCode, rightUntag, {base=rightBase, offset=rightOffset, index=rightIndex, ...}) = codeAddress(destRight, true, context) val ccRef = newCCRef() val labLess = newLabel() and labGreater = newLabel() and exitLab = newLabel() val labNotLess = newLabel() and labNotGreater = newLabel() val (lengthCode, lengthUntag, lengthArg) = codeAsUntaggedToReg(length, false (* unsigned *), context) val code = leftCode @ rightCode @ lengthCode @ leftUntag @ [BlockSimple(loadAddress{base=leftBase, offset=leftOffset, index=leftIndex, dest=vec1Reg})] @ rightUntag @ [BlockSimple(loadAddress{base=rightBase, offset=rightOffset, index=rightIndex, dest=vec2Reg})] @ lengthUntag @ [BlockSimple(CompareByteVectors{ vec1Addr=vec1Reg, vec2Addr=vec2Reg, length=lengthArg, ccRef=ccRef }), (* N.B. These are unsigned comparisons. *) BlockFlow(Conditional{ ccRef=ccRef, condition=JB, trueJump=labLess, falseJump=labNotLess }), BlockLabel labNotLess, BlockFlow(Conditional{ ccRef=ccRef, condition=JA, trueJump=labGreater, falseJump=labNotGreater }), BlockLabel labNotGreater, BlockSimple(LoadArgument{ source=IntegerConstant(tag 0), dest=mergeResult, kind=movePolyWord }), BlockFlow(Unconditional exitLab), BlockLabel labLess, BlockSimple(LoadArgument{ source=IntegerConstant(tag ~1), dest=mergeResult, kind=movePolyWord }), BlockFlow(Unconditional exitLab), BlockLabel labGreater, BlockSimple(LoadArgument{ source=IntegerConstant(tag 1), dest=mergeResult, kind=movePolyWord }), BlockLabel exitLab, BlockSimple(LoadArgument{ source=RegisterArgument mergeResult, dest=target, kind=movePolyWord })] in (code, RegisterArgument target, false) end | codeToICodeBlock({kind=BlockOpMove {isByteMove}, sourceLeft, destRight, length}, context, _, destination) = let (* Moves of 4 or 8 bytes can be done as word moves provided the alignment is correct. Although this will work for strings it is really to handle moves between SysWord and volatileRef in Foreign.Memory. Moves of 1, 2 or 3 bytes or words are converted into a sequence of byte or word moves. *) val isWordMove = case (isByteMove, length) of (true, BICConstnt(l, _)) => if not (isShort l) orelse (toShort l <> 0w4 andalso toShort l <> nativeWordSize) then NONE else let val leng = toShort l val moveKind = if toShort l = nativeWordSize then moveNativeWord else Move32Bit val isLeftAligned = case sourceLeft of {index=NONE, offset, ...} => offset mod leng = 0w0 | _ => false val isRightAligned = case destRight of {index=NONE, offset, ...} => offset mod leng = 0w0 | _ => false in if isLeftAligned andalso isRightAligned then SOME moveKind else NONE end | _ => NONE in case isWordMove of SOME moveKind => let val (leftCode, leftUntag, leftMem) = codeAddress(sourceLeft, isByteMove, context) val (rightCode, rightUntag, {base, offset, index, ...}) = codeAddress(destRight, isByteMove, context) val untaggedResReg = newUReg() val code = leftCode @ rightCode @ leftUntag @ rightUntag @ [BlockSimple(LoadArgument { source=MemoryLocation leftMem, dest=untaggedResReg, kind=moveKind}), BlockSimple(StoreArgument {source=RegisterArgument untaggedResReg, base=base, offset=offset, index=index, kind=moveKind, isMutable=true})] in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end | _ => let val vec1Reg = newUReg() and vec2Reg = newUReg() val (leftCode, leftUntag, {base=leftBase, offset=leftOffset, index=leftIndex, ...}) = codeAddress(sourceLeft, isByteMove, context) val (rightCode, rightUntag, {base=rightBase, offset=rightOffset, index=rightIndex, ...}) = codeAddress(destRight, isByteMove, context) val (lengthCode, lengthUntag, lengthArg) = codeAsUntaggedToReg(length, false (* unsigned *), context) val code = leftCode @ rightCode @ lengthCode @ leftUntag @ [BlockSimple(loadAddress{base=leftBase, offset=leftOffset, index=leftIndex, dest=vec1Reg})] @ rightUntag @ [BlockSimple(loadAddress{base=rightBase, offset=rightOffset, index=rightIndex, dest=vec2Reg})] @ lengthUntag @ [BlockSimple(BlockMove{ srcAddr=vec1Reg, destAddr=vec2Reg, length=lengthArg, isByteMove=isByteMove })] in moveIfNotAllowed(destination, code, (* Unit result *) IntegerConstant(tag 0)) end end | codeToICodeBlock({kind=BlockOpEqualByte, ...}, _, _, _) = (* TODO: Move the code from codeToICodeRev. However, that is already reversed. *) raise InternalError "codeToICodeBlock - BlockOpEqualByte" (* Already done *) and codeConditionRev(condition, context, jumpOn, jumpLabel, tailCode) = (* General case. Load the value into a register and compare it with 1 (true) *) let val ccRef = newCCRef() val (testCode, testReg) = codeToPRegRev(condition, context, tailCode) val noJumpLabel = newLabel() in BlockLabel noJumpLabel :: BlockFlow(Conditional{ccRef=ccRef, condition=if jumpOn then JE else JNE, trueJump=jumpLabel, falseJump=noJumpLabel}) :: BlockSimple(CompareLiteral{arg1=RegisterArgument testReg, arg2=tag 1, opSize=OpSize32, ccRef=ccRef}) :: testCode end (* The fixed precision functions are also used for arbitrary precision but instead of raising Overflow we need to jump to the code that handles the long format. *) and codeFixedPrecisionArith(BuiltIns.ArithAdd, arg1, BICConstnt(value, _), context, target, onOverflow) = let val ccRef = newCCRef() (* If the argument is a constant we can subtract the tag beforehand. This should always be a tagged value if the type is correct. However it's possible for it not to be if we have an arbitrary precision value. There will be a run-time check that the value is short and so this code will never be executed. It will generally be edited out by the higher level be we can't rely on that. Because it's never executed we can just put in zero. *) val constVal = if isShort value then semitag(Word.toLargeIntX(toShort value)) else 0 val (arg1Code, aReg1) = codeToPReg(arg1, context) in arg1Code @ [BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=aReg1, operand2=IntegerConstant constVal, ccRef=ccRef, opSize=polyWordOpSize})] @ onOverflow ccRef end | codeFixedPrecisionArith(BuiltIns.ArithAdd, BICConstnt(value, _), arg2, context, target, onOverflow) = let val ccRef = newCCRef() (* If the argument is a constant we can subtract the tag beforehand. Check for short - see comment above. *) val constVal = if isShort value then semitag(Word.toLargeIntX(toShort value)) else 0 val (arg2Code, aReg2) = codeToPReg(arg2, context) in arg2Code @ [BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=aReg2, operand2=IntegerConstant constVal, ccRef=ccRef, opSize=polyWordOpSize})] @ onOverflow ccRef end | codeFixedPrecisionArith(BuiltIns.ArithAdd, arg1, arg2, context, target, onOverflow) = let val aReg3 = newPReg() and ccRef = newCCRef() val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) in arg1Code @ arg2Code @ (* Subtract the tag bit from the second argument, do the addition and check for overflow. *) (* TODO: We should really do the detagging in the transform phase. It can make a better choice of the argument if one of the arguments is already untagged or if we have a constant argument. *) [BlockSimple(ArithmeticFunction{oper=SUB, resultReg=aReg3, operand1=aReg1, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize}), BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=aReg3, operand2=RegisterArgument aReg2, ccRef=ccRef, opSize=polyWordOpSize})] @ onOverflow ccRef end (* Subtraction. We can handle the special case of the second argument being a constant but not the first. *) | codeFixedPrecisionArith(BuiltIns.ArithSub, arg1, BICConstnt(value, _), context, target, onOverflow) = let val ccRef = newCCRef() (* If the argument is a constant we can subtract the tag beforehand. Check for short - see comment above. *) val constVal = if isShort value then semitag(Word.toLargeIntX(toShort value)) else 0 val (arg1Code, aReg1) = codeToPReg(arg1, context) in arg1Code @ [BlockSimple(ArithmeticFunction{oper=SUB, resultReg=target, operand1=aReg1, operand2=IntegerConstant constVal, ccRef=ccRef, opSize=polyWordOpSize})] @ onOverflow ccRef end | codeFixedPrecisionArith(BuiltIns.ArithSub, arg1, arg2, context, target, onOverflow) = let val aReg3 = newPReg() val ccRef = newCCRef() val (arg1Code, aReg1) = codeToPReg(arg1, context) val (arg2Code, aReg2) = codeToPReg(arg2, context) in arg1Code @ arg2Code @ (* Do the subtraction, test for overflow and afterwards add in the tag bit. *) [BlockSimple(ArithmeticFunction{oper=SUB, resultReg=aReg3, operand1=aReg1, operand2=RegisterArgument aReg2, ccRef=ccRef, opSize=polyWordOpSize})] @ onOverflow ccRef @ [BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=aReg3, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize})] end | codeFixedPrecisionArith(BuiltIns.ArithMult, arg1, BICConstnt(value, _), context, target, onOverflow) = let val aReg = newPReg() and argUntagged = newUReg() and resUntagged = newUReg() val mulCC = newCCRef() (* Is it better to untag the constant or the register argument? *) val constVal = if isShort value then Word.toLargeIntX(toShort value) else 0 in codeToICodeTarget(arg1, context, false, aReg) @ [BlockSimple(ArithmeticFunction{oper=SUB, resultReg=argUntagged, operand1=aReg, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize}), BlockSimple(Multiplication{resultReg=resUntagged, operand1=argUntagged, operand2=IntegerConstant constVal, ccRef=mulCC, opSize=polyWordOpSize} )] @ onOverflow mulCC @ [BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=resUntagged, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize})] end | codeFixedPrecisionArith(BuiltIns.ArithMult, BICConstnt(value, _), arg2, context, target, onOverflow) = let val aReg = newPReg() and argUntagged = newUReg() and resUntagged = newUReg() val mulCC = newCCRef() (* Is it better to untag the constant or the register argument? *) val constVal = if isShort value then Word.toLargeIntX(toShort value) else 0 in codeToICodeTarget(arg2, context, false, aReg) @ [BlockSimple(ArithmeticFunction{oper=SUB, resultReg=argUntagged, operand1=aReg, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize}), BlockSimple(Multiplication{resultReg=resUntagged, operand1=argUntagged, operand2=IntegerConstant constVal, ccRef=mulCC, opSize=polyWordOpSize} )] @ onOverflow mulCC @ [BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=resUntagged, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize})] end | codeFixedPrecisionArith(BuiltIns.ArithMult, arg1, arg2, context, target, onOverflow) = let val aReg1 = newPReg() and aReg2 = newPReg() and arg1Untagged = newUReg() and arg2Untagged = newUReg() and resUntagged = newUReg() val mulCC = newCCRef() (* This is almost the same as the word operation except we use a signed shift and check for overflow. *) in codeToICodeTarget(arg1, context, false, aReg1) @ codeToICodeTarget(arg2, context, false, aReg2) @ (* Shift one argument and subtract the tag from the other. It's possible this could be reordered if we have a value that is already untagged. *) [BlockSimple(UntagValue{source=aReg1, dest=arg1Untagged, isSigned=true (* Signed shift here. *), cache=NONE, opSize=polyWordOpSize}), BlockSimple(ArithmeticFunction{oper=SUB, resultReg=arg2Untagged, operand1=aReg2, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize}), BlockSimple(Multiplication{resultReg=resUntagged, operand1=arg1Untagged, operand2=RegisterArgument arg2Untagged, ccRef=mulCC, opSize=polyWordOpSize} )] @ onOverflow mulCC @ [BlockSimple(ArithmeticFunction{oper=ADD, resultReg=target, operand1=resUntagged, operand2=IntegerConstant 1, ccRef=newCCRef(), opSize=polyWordOpSize})] end | codeFixedPrecisionArith(BuiltIns.ArithQuot, arg1, arg2, context, target, _) = let val aReg1 = newPReg() and aReg2 = newPReg() val arg1Untagged = newUReg() and arg2Untagged = newUReg() val quotient = newUReg() and remainder = newUReg() in codeToICodeTarget(arg1, context, false, aReg1) @ codeToICodeTarget(arg2, context, false, aReg2) @ (* Shift both of the arguments to remove the tags. We don't test for zero here - that's done explicitly. *) [BlockSimple(UntagValue{source=aReg1, dest=arg1Untagged, isSigned=true, cache=NONE, opSize=polyWordOpSize}), BlockSimple(UntagValue{source=aReg2, dest=arg2Untagged, isSigned=true, cache=NONE, opSize=polyWordOpSize}), BlockSimple(Division { isSigned = true, dividend=arg1Untagged, divisor=RegisterArgument arg2Untagged, quotient=quotient, remainder=remainder, opSize=polyWordOpSize }), BlockSimple(TagValue { source=quotient, dest=target, isSigned=true, opSize=polyWordOpSize})] end | codeFixedPrecisionArith(BuiltIns.ArithRem, arg1, arg2, context, target, _) = let (* Identical to Quot except that the result is the remainder. *) val aReg1 = newPReg() and aReg2 = newPReg() val arg1Untagged = newUReg() and arg2Untagged = newUReg() val quotient = newUReg() and remainder = newUReg() in codeToICodeTarget(arg1, context, false, aReg1) @ codeToICodeTarget(arg2, context, false, aReg2) @ (* Shift both of the arguments to remove the tags. *) [BlockSimple(UntagValue{source=aReg1, dest=arg1Untagged, isSigned=true, cache=NONE, opSize=polyWordOpSize}), BlockSimple(UntagValue{source=aReg2, dest=arg2Untagged, isSigned=true, cache=NONE, opSize=polyWordOpSize}), BlockSimple(Division { isSigned = true, dividend=arg1Untagged, divisor=RegisterArgument arg2Untagged, quotient=quotient, remainder=remainder, opSize=polyWordOpSize }), BlockSimple(TagValue { source=remainder, dest=target, isSigned=true, opSize=polyWordOpSize})] end | codeFixedPrecisionArith(_, _, _, _, _, _) = raise InternalError "codeToICode: FixedPrecisionArith - unimplemented operation" (* Generate code for floating point arguments where one of the arguments must be in a register. If the first argument is in a register use that, if the second is in a register and it's commutative use that otherwise load the first argument into a register. *) and codeFPBinaryArgsRev(arg1, arg2, precision, commutative, context, tailCode) = let val (arg1Code, arg1Value) = codeFPArgument(arg1, precision, context, tailCode) val (arg2Code, arg2Value) = codeFPArgument(arg2, precision, context, arg1Code) in case (arg1Value, arg2Value, commutative) of (RegisterArgument fpReg, _, _) => (arg2Code, fpReg, arg2Value) | (_, RegisterArgument fpReg, Commutative) => (arg2Code, fpReg, arg1Value) | (arg1Val, _, _) => let val fpReg = newUReg() val moveOp = case precision of BuiltIns.PrecDouble => MoveDouble | BuiltIns.PrecSingle => MoveFloat in (BlockSimple(LoadArgument{source=arg1Val, dest=fpReg, kind=moveOp}) :: arg2Code, fpReg, arg2Value) end end (* Generate code to evaluate a floating point argument. The aim of this code is to avoid the overhead of untagging a short-precision floating point value in memory. *) and codeFPArgument(BICConstnt(value, _), _, _, tailCode) = let val argVal = (* Single precision constants in 64-bit mode are represented by the value shifted left 32 bits. A word is shifted left one bit so the result is 0w31. *) if isShort value then IntegerConstant(Word.toLargeInt(Word.>>(toShort value, 0w31))) else AddressConstant value in (tailCode, argVal) end | codeFPArgument(arg, precision, context, tailCode) = ( case (precision, wordSize) of (BuiltIns.PrecSingle, 0w8) => (* If this is a single precision value and the word size is 8 the values are tagged. If it is memory we can load the value directly from the high-order word. *) let val memOrReg = { anyConstant=false, const32s=false, memAddr=true, existingPreg=true } val (code, result, _) = codeToICodeRev(arg, context, false, Allowed memOrReg, tailCode) in case result of RegisterArgument argReg => let val fpReg = newUReg() in (BlockSimple(UntagFloat{source=RegisterArgument argReg, dest=fpReg, cache=NONE}) :: code, RegisterArgument fpReg) end | MemoryLocation{offset, base, index, ...} => (code, MemoryLocation{offset=offset+4, base=base, index=index, cache=NONE}) | _ => raise InternalError "codeFPArgument" end | _ => (* Otherwise the value is boxed. *) let val (argCode, argReg) = codeToPRegRev(arg, context, tailCode) in (argCode, wordAt argReg) end ) (* Code an address. The index is optional. *) and codeAddressRev({base, index=SOME index, offset}, true (* byte move *), context, tailCode) = let (* Byte address with index. The index needs to be untagged. *) val indexReg1 = newUReg() val (codeBase, baseReg) = codeToPRegRev(base, context, tailCode) val (codeIndex, indexReg) = codeToPRegRev(index, context, codeBase) val untagCode = [BlockSimple(UntagValue{dest=indexReg1, source=indexReg, isSigned=false, cache=NONE, opSize=polyWordOpSize})] val (codeLoadAddr, realBase) = if targetArch = ObjectId32Bit then let val addrReg = newUReg() in ([BlockSimple(LoadEffectiveAddress{ base=SOME baseReg, offset=0, index=ObjectIndex, dest=addrReg, opSize=nativeWordOpSize})], addrReg) end else ([], baseReg) val memResult = {base=realBase, offset=Word.toInt offset, index=MemIndex1 indexReg1, cache=NONE} in (codeLoadAddr @ codeIndex, untagCode, memResult) end | codeAddressRev({base, index=SOME index, offset}, false (* word move *), context, tailCode) = let (* Word address with index. We can avoid untagging the index by adjusting the multiplier and offset *) val (codeBase, baseReg) = codeToPRegRev(base, context, tailCode) val (codeIndex, indexReg) = codeToPRegRev(index, context, codeBase) val (codeLoadAddr, realBase) = if targetArch = ObjectId32Bit then let val addrReg = newUReg() in ([BlockSimple(LoadEffectiveAddress{ base=SOME baseReg, offset=0, index=ObjectIndex, dest=addrReg, opSize=nativeWordOpSize})], addrReg) end else ([], baseReg) val iOffset = Word.toInt offset handle Overflow => 0 (* See below: special case may not happen. *) val memResult = if wordSize = 0w8 then {base=realBase, offset=iOffset-4, index=MemIndex4 indexReg, cache=NONE} else {base=realBase, offset=iOffset-2, index=MemIndex2 indexReg, cache=NONE} in (codeLoadAddr @ codeIndex, [], memResult) end | codeAddressRev({base, index=NONE, offset}, _, context, tailCode) = let val (codeBase, baseReg) = codeToPRegRev(base, context, tailCode) (* A negative value for "offset" will produce an overflow at compile time. It should never be reached at run-time because of bounds checking. See Test192. *) val iOffset = Word.toInt offset handle Overflow => 0 val memResult = {offset=iOffset, base=baseReg, index=memIndexOrObject, cache=NONE} in (codeBase, [], memResult) end and codeAddress(addr, isByte, context) = let val (code, untag, res) = codeAddressRev(addr, isByte, context, []) in (List.rev code, untag, res) end (* C-memory operations are slightly different. The base address is a LargeWord.word value. The index is a byte index so may have to be untagged. *) and codeCAddress({base, index=SOME index, offset}, 0w1, context) = let (* Byte address with index. The index needs to be untagged. *) val untaggedBaseReg = newUReg() and indexReg1 = newUReg() val (codeBase, baseReg) = codeToPReg(base, context) and (codeIndex, indexReg) = codeToPReg(index, context) val untagCode = [BlockSimple(LoadArgument{source=wordAt baseReg, dest=untaggedBaseReg, kind=moveNativeWord}), BlockSimple(UntagValue{dest=indexReg1, source=indexReg, isSigned=false, cache=NONE, opSize=polyWordOpSize})] val memResult = {base=untaggedBaseReg, offset=Word.toInt offset, index=MemIndex1 indexReg1, cache=NONE} in (codeBase @ codeIndex, untagCode, memResult) end | codeCAddress({base, index=SOME index, offset}, size, context) = let (* Non-byte address with index. By using an appropriate multiplier we can avoid having to untag the index. *) val untaggedBaseReg = newUReg() val (codeBase, baseReg) = codeToPReg(base, context) and (codeIndex, indexReg) = codeToPReg(index, context) val untagCode = [BlockSimple(LoadArgument{source=wordAt baseReg, dest=untaggedBaseReg, kind=moveNativeWord})] val memResult = case size of 0w2 => {base=untaggedBaseReg, offset=Word.toInt offset-1, index=MemIndex1 indexReg, cache=NONE} | 0w4 => {base=untaggedBaseReg, offset=Word.toInt offset-2, index=MemIndex2 indexReg, cache=NONE} | 0w8 => {base=untaggedBaseReg, offset=Word.toInt offset-4, index=MemIndex4 indexReg, cache=NONE} | _ => raise InternalError "codeCAddress: unknown size" in (codeBase @ codeIndex, untagCode, memResult) end | codeCAddress({base, index=NONE, offset}, _, context) = let val untaggedBaseReg = newUReg() val (codeBase, baseReg) = codeToPReg(base, context) val untagCode = [BlockSimple(LoadArgument{source=wordAt baseReg, dest=untaggedBaseReg, kind=moveNativeWord})] val memResult = {offset=Word.toInt offset, base=untaggedBaseReg, index=NoMemIndex, cache=NONE} in (codeBase, untagCode, memResult) end (* Return an untagged value. If we have a constant just return it. Otherwise return the code to evaluate the argument, the code to untag it and the reference to the untagged register. *) and codeAsUntaggedToRegRev(BICConstnt(value, _), isSigned, _, tailCode) = let (* Should always be short except for unreachable code. *) val untagReg = newUReg() val cval = if isShort value then toShort value else 0w0 val cArg = IntegerConstant(if isSigned then Word.toLargeIntX cval else Word.toLargeInt cval) (* Don't tag *) val untag = [BlockSimple(LoadArgument{source=cArg, dest=untagReg, kind=movePolyWord})] in (tailCode, untag, untagReg) (* Don't tag. *) end | codeAsUntaggedToRegRev(arg, isSigned, context, tailCode) = let val untagReg = newUReg() val (code, srcReg) = codeToPRegRev(arg, context, tailCode) val untag = [BlockSimple(UntagValue{source=srcReg, dest=untagReg, isSigned=isSigned, cache=NONE, opSize=polyWordOpSize})] in (code, untag, untagReg) end and codeAsUntaggedToReg(arg, isSigned, context) = let val (code, untag, untagReg) = codeAsUntaggedToRegRev(arg, isSigned, context, []) in (List.rev code, untag, untagReg) end (* Return the argument as an untagged value. We separate evaluating the argument from untagging because we may have to evaluate other arguments and that could involve a function call and we can't save the value to the stack after we've untagged it. Currently this is only used for byte values but we may have to be careful if we use it for word values on the X86. Moving an untagged value into a register might look like loading a constant address. *) and codeAsUntaggedByte(BICConstnt(value, _), isSigned, _) = let val cval = if isShort value then toShort value else 0w0 val cArg = IntegerConstant(if isSigned then Word.toLargeIntX cval else Word.toLargeInt cval) (* Don't tag *) in ([], [], cArg) end | codeAsUntaggedByte(arg, isSigned, context) = let val untagReg = newUReg() val (code, argReg) = codeToPReg(arg, context) val untag = [BlockSimple(UntagValue{source=argReg, dest=untagReg, isSigned=isSigned, cache=NONE, opSize=OpSize32})] in (code, untag, RegisterArgument untagReg) end (* Allocate memory. This is used both for true variable length cells and also for longer constant length cells. *) and allocateMemoryVariable(numWords, flags, initial, context, destination) = let val target = asTarget destination (* With the exception of flagReg all these registers are modified by the code. So, we have to copy the size value into a new register. *) val sizeReg = newPReg() and initReg = newPReg() val sizeReg2 = newPReg() val untagSizeReg = newUReg() and initAddrReg = newPReg() and allocReg = newPReg() val sizeCode = codeToICodeTarget(numWords, context, false, sizeReg) and (flagsCode, flagUntag, flagArg) = codeAsUntaggedByte(flags, false, context) (* We're better off deferring the initialiser if possible. If the value is a constant we don't have to save it. *) val (initCode, initResult, _) = codeToICode(initial, context, false, Allowed allowDefer) in (sizeCode @ flagsCode @ initCode @ [(* We need to copy the size here because AllocateMemoryVariable modifies the size in order to store the length word. This is unfortunate especially as we're going to untag it anyway. *) BlockSimple(LoadArgument{source=RegisterArgument sizeReg, dest=sizeReg2, kind=movePolyWord}), BlockSimple(AllocateMemoryVariable{size=sizeReg, dest=allocReg, saveRegs=[]})] @ flagUntag @ [BlockSimple(StoreArgument{ source=flagArg, base=allocReg, offset= ~1, index=memIndexOrObject, kind=MoveByte, isMutable=false}), (* We need to copy the address here because InitialiseMem modifies all its arguments. *) BlockSimple( if targetArch = ObjectId32Bit then LoadEffectiveAddress{ base=SOME allocReg, offset=0, index=ObjectIndex, dest=initAddrReg, opSize=nativeWordOpSize} else LoadArgument{source=RegisterArgument allocReg, dest=initAddrReg, kind=movePolyWord}), BlockSimple(UntagValue{source=sizeReg2, dest=untagSizeReg, isSigned=false, cache=NONE, opSize=polyWordOpSize}), BlockSimple(LoadArgument{source=initResult, dest=initReg, kind=movePolyWord}), BlockSimple(InitialiseMem{size=untagSizeReg, init=initReg, addr=initAddrReg}), BlockSimple InitialisationComplete, BlockSimple(LoadArgument{source=RegisterArgument allocReg, dest=target, kind=movePolyWord})], RegisterArgument target, false) end (*Turn the codetree structure into icode. *) val bodyContext = {loopArgs=NONE, stackPtr=0, currHandler=NONE, overflowBlock=ref NONE} val (bodyCode, _, bodyExited) = codeToICodeRev(body, bodyContext, true, SpecificPReg resultTarget, beginInstructions) val icode = if bodyExited then bodyCode else returnInstruction(bodyContext, resultTarget, bodyCode) (* Turn the icode list into basic blocks. The input list is in reverse so as part of this we reverse the list. *) local val resArray = Array.array(!labelCounter, BasicBlock{ block=[], flow=ExitCode }) fun createEntry (blockNo, block, flow) = Array.update(resArray, blockNo, BasicBlock{ block=block, flow=flow}) fun splitCode([], _, _) = (* End of code. We should have had a BeginFunction. *) raise InternalError "splitCode - no begin" | splitCode(BlockBegin args :: _, sinceLabel, flow) = (* Final instruction. Create the initial block and exit. *) createEntry(0, BeginFunction args ::sinceLabel, flow) | splitCode(BlockSimple instr :: rest, sinceLabel, flow) = splitCode(rest, instr :: sinceLabel, flow) | splitCode(BlockLabel label :: rest, sinceLabel, flow) = (* Label - finish this block and start another. *) ( createEntry(label, sinceLabel, flow); (* Default to a jump to this label. That is used if we have assumed a drop-through. *) splitCode(rest, [], Unconditional label) ) | splitCode(BlockExit instr :: rest, _, _) = splitCode(rest, [instr], ExitCode) | splitCode(BlockFlow flow :: rest, _, _) = splitCode(rest, [], flow) | splitCode(BlockRaiseAndHandle(instr, handler) :: rest, _, _) = splitCode(rest, [instr], UnconditionalHandle handler) | splitCode(BlockOptionalHandle{call, handler, label} :: rest, sinceLabel, flow) = let (* A function call within a handler. This could go to the handler but if there is no exception will go to the next instruction. Also includes JumpLoop since the stack check could result in an Interrupt exception. *) in createEntry(label, sinceLabel, flow); splitCode(rest, [call], ConditionalHandle{handler=handler, continue=label}) end in val () = splitCode(icode, [], ExitCode) val resultVector = Array.vector resArray end open ICODETRANSFORM val pregProperties = Vector.fromList(List.rev(! pregPropList)) in codeICodeFunctionToX86{blocks = resultVector, functionName = name, pregProps = pregProperties, ccCount= ! ccRefCounter, debugSwitches = debugSwitches, resultClosure = resultClosure} end fun gencodeLambda(lambda, debugSwitches, closure) = let open DEBUG Universal (*val debugSwitches = [tagInject Pretty.compilerOutputTag (Pretty.prettyPrint(print, 70)), tagInject assemblyCodeTag true] @ debugSwitches*) in codeFunctionToX86(lambda, debugSwitches, closure) end structure Foreign = X86FOREIGN structure Sharing = struct type backendIC = backendIC and bicLoadForm = bicLoadForm and argumentType = argumentType and closureRef = closureRef end end; diff --git a/mlsource/MLCompiler/DATATYPE_REP.ML b/mlsource/MLCompiler/DATATYPE_REP.ML index 280d48cd..c1da502a 100644 --- a/mlsource/MLCompiler/DATATYPE_REP.ML +++ b/mlsource/MLCompiler/DATATYPE_REP.ML @@ -1,685 +1,683 @@ (* - Copyright (c) 2009, 2013, 2015-16 David C.J. Matthews + Copyright (c) 2009, 2013, 2015-16, 2020 David C.J. Matthews This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) (* Title: Operations on global and local values. Author: Dave Matthews, Cambridge University Computer Laboratory Copyright Cambridge University 1986 *) functor DATATYPE_REP ( structure CODETREE : CODETREESIG structure STRUCTVALS : STRUCTVALSIG; structure TYPESTRUCT : TYPETREESIG structure MISC : sig exception InternalError of string; (* compiler error *) val quickSort : ('a -> 'a -> bool) -> 'a list -> 'a list end; structure ADDRESS : AddressSig structure COPIER: COPIERSIG structure TYPEIDCODE: TYPEIDCODESIG sharing STRUCTVALS.Sharing = TYPESTRUCT.Sharing = COPIER.Sharing = CODETREE.Sharing = ADDRESS = MISC = TYPEIDCODE.Sharing ) : DATATYPEREPSIG = struct open MISC; open CODETREE; open TYPESTRUCT; (* Open this first because unitType is in STRUCTVALS as well. *) open Universal; (* for tag etc. *) open STRUCTVALS; open ADDRESS; open TYPEIDCODE open COPIER val length = List.length; val orb = Word8.orb infix 7 orb; (* These are the possible representations of a value constructor. *) datatype representations = RefForm (* As for OnlyOne but must be a monotype. *) | UnitForm (* If the only value in an enumeration. *) | OnlyOne (* If only one constructor, there is no tag or box. *) (* Could be replaced by "UnboxedForm"? *) | EnumForm of { tag: word, maxTag: word } (* Enumeration - argument is the number. *) | ShortForm of word (* As EnumForm except that one datatype is BoxedForm. *) | BoxedForm (* Boxed but not tagged (only unary constructor) *) | UnboxedForm of int (* Unboxed and untagged (only unary constructor) *) | ConstForm of { value: machineWord, maxTag: word} (* Constant - argument is a tagged value. *) | TaggedBox of { tag: word, maxTag: word } (* Union - tagged and boxed. i.e. the representation is a pair whose first word is the tag and second is the value. *) | TaggedTuple of { tag: word, maxTag: word, size: int } (* Union - tagged but with in-line tuple. i.e. for a tuple of size n the representation is a tuple of size n+1 whose first word contains the tag. *) val arg1 = mkLoadArgument 0 (* saves a lot of garbage *) val mutableFlags = F_words orb F_mutable; (* allocate 1 mutable word, initialise to "v" *) fun refApplyCode v = mkAllocateWordMemory(mkConst (toMachineWord 1), mkConst (toMachineWord mutableFlags), v) local fun mkTag (tag:word) : codetree = mkConst (toMachineWord tag); (* How to apply the constructor at run-time or, if, it's a constant make it now. *) fun constrApply (test: representations, arg) : codetree = let fun tagTupleApplyCode (tag, n, arg) = mkEval (mkInlproc (mkDatatype (mkTag tag :: List.tabulate(n, fn i => mkInd(i, arg1))), 1, "", [], 0), [arg]) (* Even though unboxed tuples (e.g. list cells) are the same as tuples we need to add this extra step so that the result is a variant tuple i.e. the optimiser can tell that this may not always be a tuple. *) fun tupleApplyCode (n, arg) = mkEval (mkInlproc (mkDatatype (List.tabulate(n, fn i => mkInd(i, arg1))), 1, "", [], 0), [arg]) in case test of UnboxedForm 0 => arg (* Function - never detupled. *) | UnboxedForm size => tupleApplyCode(size, arg) | BoxedForm => mkDatatype [arg] | RefForm => refApplyCode arg | TaggedBox{tag, ...} => mkDatatype [mkTag tag, arg] | TaggedTuple{tag, size, ...} => tagTupleApplyCode(tag, size, arg) | OnlyOne => arg | ConstForm{ value, ...} => mkConst value (* tagged value. *) | EnumForm{tag, ...} => mkConst (toMachineWord tag) | ShortForm tag => mkConst (toMachineWord tag) | UnitForm => CodeZero end (* The run-time test whether a value matches a constructor. *) fun constrMatch (test: representations, value:codetree) : codetree = let - fun testTag (tag, v) = mkEqualWord(mkTag tag, v) - fun testTagOf(tag, maxTag, v) = mkTagTest (v, tag, maxTag) fun testBoxedTagOf (tag, maxTag, v) = testTagOf (tag, maxTag, mkInd (0, v)) (* Tag is first field. It is always present and is always the tag so we can use mkInd here rather than mkVarField. *) val testBoxed = mkNot o mkIsShort (* not (isShort v) *) (* get the tag from a TaggedBox or ConstForm *) fun loadTag (u: machineWord) : machineWord = loadWord (toAddress u, 0w0); (* tag is first field *) in case test of UnboxedForm _ => testBoxed value | BoxedForm => testBoxed value | RefForm => CodeTrue | EnumForm{tag, maxTag} => testTagOf(tag, maxTag, value) - | ShortForm tag => testTag(tag, value) + | ShortForm tag => mkEqualPointerOrWord(mkTag tag, value) (* Could be an address. *) | TaggedBox{tag, maxTag} => testBoxedTagOf(tag, maxTag, value) | TaggedTuple{tag, maxTag, ...} => testBoxedTagOf(tag, maxTag, value) | ConstForm{value=c, maxTag} => testBoxedTagOf(toShort (loadTag c), maxTag, value) | OnlyOne => CodeTrue | UnitForm => CodeTrue end (* The run-time code to destruct a construction. *) (* shouldn't the CodeZero's raise an exception instead? No, because I think there are circumstances in which the destructor code is created even for nullary constructors. *) fun constrDestruct (test: representations, value: codetree) : codetree = let (* Copy out the fields and build a tuple. Used either if we have a tagged tuple (offset 1) or a tuple that does not need boxing (because we need to use mkVarField to extract the fields). *) fun tupleDestructCode (n, arg, offset) = mkEval ( mkInlproc (mkTuple (List.tabulate(n, fn i => mkVarField(i+offset, arg1))), 1, "", [], 0), [arg]) (* Use loadWord not indirect because the optimiser reorders indirections. *) fun refDestructCode v = mkLoadOperation(LoadStoreMLWord{isImmutable=false}, v, CodeZero) in case test of UnboxedForm 0 => value (* Function - never detupled. *) | UnboxedForm size => tupleDestructCode(size, value, 0) | BoxedForm => mkVarField (0, value) | RefForm => refDestructCode value | TaggedBox _ => mkVarField (1, value) (* contents is second field of record *) | TaggedTuple { size, ...} => tupleDestructCode(size, value, 1) | OnlyOne => value | EnumForm _ => CodeZero (* To keep optimiser happy. *) | ShortForm _ => CodeZero (* To keep optimiser happy. *) | ConstForm _ => CodeZero (* (rather than raising an exception) *) | UnitForm => CodeZero end open ValueConstructor in (* Constructors are now represented as run-time values. A nullary constructor is a pair consisting of a test function and the constructor value. A unary constructor is a triple: a test function, an injection function and a projection function. The above applies to monotypes. If this is a polytype each of these is actually a function from the base type values to the functions. *) fun createNullaryConstructor (test, tvs, name): codetree = let val numTypes = if justForEqualityTypes then 0 else List.length tvs val testFn = mkInlproc(constrMatch(test, arg1), 1, name, [], 0) (* Test function. *) and constrVal = constrApply(test, CodeZero) (* Value. *) in if numTypes = 0 then createNullaryConstr{ testMatch = testFn, constrValue = constrVal } else createNullaryConstr{ testMatch = mkInlproc(testFn, numTypes, name, [], 0), constrValue = mkInlproc(constrVal, numTypes, name, [], 0)} end fun createUnaryConstructor(test: representations, tvs, name: string): codetree = let val numTypes = if justForEqualityTypes then 0 else List.length tvs val testMatch = mkInlproc(constrMatch(test, arg1), 1, name, [], 0) (* Test function. *) and injectValue = mkInlproc(constrApply(test, arg1), 1, name, [], 0) (* Injection function. *) and projectValue = mkInlproc(constrDestruct(test, arg1), 1, name, [], 0) (* Projection function. *) in if numTypes = 0 then createValueConstr{testMatch = testMatch, injectValue = injectValue, projectValue = projectValue } else createValueConstr{ testMatch = mkInlproc(testMatch, numTypes, name, [], 0), injectValue = mkInlproc(injectValue, numTypes, name, [], 0), projectValue = mkInlproc(projectValue, numTypes, name, [], 0)} end end (* RefForm is used for "ref" (only). We use various representations of datatype / abstype constructors. Nullary constructors are represented as: UnitForm (if it's the only constructor in the datatype) EnumForm (if all the constructors are nullary) ShortForm (if there's one non-nullary constructor) ConstForm (otherwise) Unary constructors are represented as: OnlyOne (if it's the only constructor in the datatype) UnboxedForm (if it's the only unary constructor and the argument is always a tuple e.g. list) TaggedTuple (if it's not the only unary constructor, applied to a tuple) BoxedForm (if it's the only unary constructor and the argument may not be a tuple e.g. SOME) TaggedBox (otherwise) Note that we use ConstForm, not EnumForm, for nullary constructors when the unary constructors are represented as TaggedTuple/TaggedBox because that allows the TaggedBox test to be: fn w => wordEq (loadWord (w,0), tag) rather than: fn w => not (isShort w) andalso wordEq (loadWord (w,0), tag) Note that EnumForm and ShortForm differ in that the tests for EnumForm use mkTagTest which can be converted into an indexed case. This can't be used for ShortForm because the values for the datatype include addresses. *) datatype constructorKind = Nullary (* a nullary constructor *) | UnaryGeneric (* a normal unary constructor *) | UnaryFunction (* unary constructor applied to a function *) | UnaryTupled of int (* a unary constructor applied to a tuple of size n *) ; fun getTupleKind t = case t of (* We cannot have flexible records here. All the fields must be listed. *) LabelledType {recList = [{typeof=t', ...}], ...} => (* Singleton records are always represented simply by the value. *) getTupleKind t' | LabelledType {recList, ...} => UnaryTupled (length recList) | FunctionType _ => UnaryFunction | TypeConstruction {constr, args, ...} => ( (* We may have a type equivalence or this may be a datatype. *) if tcIsAbbreviation constr then getTupleKind (makeEquivalent(constr, args)) else if sameTypeId (tcIdentifier constr, tcIdentifier refConstr) then UnaryGeneric (* A tuple ref is NOT the same as the tuple. *) else (* Datatype. For the moment we only consider datatypes with a single constructor since we want to find the width of the tuple. At present we simply return UnaryGeneric for all other cases but it might be helpful to return a special result when we have a datatype which we know will always be boxed. *) (* case tcConstructors constr of [Value{typeOf, class=Constructor{nullary=false, ...}, ...}] => (* This may be a polymorphic datatype in which case we have to invert the constructor to find the base type. e.g. we may have an instance (int*int) t where t was declared as datatype 'a t = X of 'a .*) getTupleKind(constructorResult(typeOf, args)) | _ => UnaryGeneric *) UnaryGeneric ) | _ => UnaryGeneric (* This now creates the functions as well as choosing the representation. *) (* N.B. The representation for the "context" and "pretty" datatypes is defined in Pretty.sml. Any changes here that may affect the representation of a datatype may require changes there as well. *) fun chooseConstrRepr(cs, tvs: types list) = let fun checkArgKind (name, EmptyType) = (Nullary, name) | checkArgKind (name, argType) = (getTupleKind argType, name) val kinds = map checkArgKind cs; fun chooseRepr [(Nullary, name)] = [createNullaryConstructor(UnitForm, tvs, name)] | chooseRepr [(UnaryGeneric, name)] = [createUnaryConstructor(OnlyOne, tvs, name)] | chooseRepr [(UnaryFunction, name)] = [createUnaryConstructor(OnlyOne, tvs, name)] | chooseRepr [(UnaryTupled _, name)] = [createUnaryConstructor(OnlyOne, tvs, name)] | chooseRepr l = let val unaryCount = List.foldl(fn((Nullary, _), n) => n | (_,n) => n+1) 0 l in case unaryCount of 0 => (* All are nullary. *) let val maxTag = Word.fromInt(List.length l)-0w1 (* Largest no is length-1 *) fun createRepr(_, []) = [] | createRepr(n, (_, name) :: t) = createNullaryConstructor(EnumForm{tag=n, maxTag=maxTag}, tvs, name) :: createRepr (n + 0w1, t) in createRepr(0w0, l) end | 1 => let (* We use this version if all the constructors are nullary (i.e. constants) except one. The unary constructor is represented by the boxed value and the nullary constructors by untagged integers. *) (* Note that "UnaryTupled 0" (which would arise as a result of a declaration of the form datatype t = A of () | ... ) can't be represented as "UnboxedForm" because "{}" is represented as a short (unboxed) integer. *) fun chooseOptimisedRepr1(_, _, []) = [] | chooseOptimisedRepr1(n, tvs, (Nullary, name) :: t) = createNullaryConstructor(ShortForm n, tvs, name) :: chooseOptimisedRepr1 (n + 0w1, tvs, t) | chooseOptimisedRepr1(n, tvs, (UnaryGeneric, name) :: t) = createUnaryConstructor(BoxedForm, tvs, name) :: chooseOptimisedRepr1(n, tvs, t) | chooseOptimisedRepr1(n, tvs, (UnaryFunction, name) :: t) = createUnaryConstructor(UnboxedForm 0, tvs, name) :: chooseOptimisedRepr1(n, tvs, t) | chooseOptimisedRepr1(n, tvs, (UnaryTupled 0, name) :: t) = createUnaryConstructor(BoxedForm, tvs, name) :: chooseOptimisedRepr1(n, tvs, t) | chooseOptimisedRepr1(n, tvs, (UnaryTupled s, name) :: t) = createUnaryConstructor(UnboxedForm s, tvs, name) :: chooseOptimisedRepr1(n, tvs, t) in chooseOptimisedRepr1(0w0, tvs, l) (* can save the box *) end | _ => let (* We use this version there's more than 1 unary constructor. *) (* With this representation constructors of small tuples make tuples of size n+1 whose first word is the tag. Nullary constructors are represented by single word objects containing the tag. *) val maxTag = Word.fromInt(List.length l) - 0w1 (* Largest no is length - 1 *) fun chooseOptimisedRepr2(_, _, []) = [] | chooseOptimisedRepr2(n, tvs, h :: t) = let val repr = case h of (Nullary, name) => let (* Make an object with the appropriate tag. Doing it here means we only do it once for this object. *) fun genConstForm (n :word) : representations = let val vec : address = allocWordData (0w1, F_words, toMachineWord n) (* The new call does not require locking but the old code still sets the F_mutable bit. *) val _ = if isMutable vec then lock vec else () in ConstForm{value=toMachineWord vec, maxTag=maxTag} end in createNullaryConstructor(genConstForm n, tvs, name) end | (UnaryGeneric, name) => createUnaryConstructor(TaggedBox{tag=n, maxTag=maxTag}, tvs, name) | (UnaryFunction, name) => createUnaryConstructor(TaggedBox{tag=n, maxTag=maxTag}, tvs, name) | (UnaryTupled i, name) => createUnaryConstructor( if i <= 4 (*!maxPacking*) then TaggedTuple {tag=n, size=i, maxTag=maxTag} else TaggedBox{tag=n, maxTag=maxTag}, tvs, name) in repr :: chooseOptimisedRepr2(n + 0w1, tvs, t) end; in chooseOptimisedRepr2(0w0, tvs, l) (* can use tagged tuples *) end end; fun makeFun c = mkInlproc(c, List.length tvs, "boxed/size", [], 0) val (boxed, size) = case tvs of [] => (* Monotype *) (TypeValue.boxedEither, TypeValue.singleWord) | _ => (* Polytype *) (makeFun TypeValue.boxedEither, makeFun TypeValue.singleWord) in { constrs = chooseRepr kinds, boxed = boxed, size = size } end; (* RefForm, NilForm and ConsForm are only used for built-in types *) (*****************************************************************************) (* Standard values and exceptions. *) (*****************************************************************************) (* Build a datatype within the basis. *) fun buildBasisDatatype(tcName, tIdPath, tyVars, isEqType: bool, mkValConstrs: typeConstrs -> (values * codetree) list * codetree * codetree) = let (* Create a temporary datatype. The "name" we put in here is usually the same as the type constructor name except for datatypes in the PolyML structure which have the PolyML prefix. *) val arity = List.length tyVars val description = basisDescription tIdPath val id = makeBoundId(arity, Local{addr = ref ~1, level = ref baseLevel}, 0 (* IdNumber*), isEqType, true, description) val dtype = makeTypeConstructor (tcName, tyVars, id, [DeclaredAt inBasis]); (* Build the constructors. *) val (valConstrsAndDecs, boxedCode, sizeCode) = mkValConstrs dtype (* The constructors have to be ordered as in genValueConstrs in PARSE_TREE. *) fun leq (Value{name=xname, ...}, _) (Value{name=yname, ...}, _) = xname < yname; val sortedConstrs = quickSort leq valConstrsAndDecs; val initialTypeSet = TypeConstrSet(dtype, (List.map #1 valConstrsAndDecs)) val addrs = ref 0 fun mkAddrs n = ! addrs before (addrs := !addrs+n) fun declConstr(Value{access=Local{addr, level}, ...}, repr) = let val newAddr = mkAddrs 1 in addr := newAddr; level := baseLevel; (mkDec(newAddr, repr), mkLoadLocal newAddr) end | declConstr _ = raise InternalError "declConstr: not local" val (declConstrs, loadConstrs) = ListPair.unzip(List.map declConstr sortedConstrs) val defMap = TypeVarMap.defaultTypeVarMap(mkAddrs, baseLevel) (* Create the datatype. Sets the address of the local in "id". *) val dtCode = createDatatypeFunctions( [{typeConstr=initialTypeSet, eqStatus=isEqType, boxedCode=boxedCode, sizeCode=sizeCode}], mkAddrs, baseLevel, defMap, true) (* Compile and execute the code to build the functions and extract the result. *) val globalCode = genCode( mkEnv( declConstrs @ TypeVarMap.getCachedTypeValues defMap @ dtCode, mkTuple(codeId(id, baseLevel) :: loadConstrs)), [], !addrs)() val newId = makeFreeId(arity, Global(mkInd(0, globalCode)), isEqType, description) (* Get the value constructors out as globals. *) fun mkGlobal((Value{name, typeOf, class, locations, ...}, _), (decs, offset)) = (decs @ [Value{name=name, typeOf=typeOf, class=class, locations=locations, references=NONE, instanceTypes=NONE, access=Global(mkInd(offset, globalCode))}], offset+1) val (gConstrs, _) = List.foldl mkGlobal ([], 1 (* Offset 0 is the type ID *)) sortedConstrs (* Finally copy the datatype to put in the code. *) in fullCopyDatatype(TypeConstrSet(dtype, gConstrs), fn 0 => newId | _ => raise Subscript, "") end (* Nil and :: are used in parsetree for lists constructed using [ ... ] and are also used for initialisation. *) local fun makeConsAndNil listType = let val listTypeVars = tcTypeVars listType; val alpha = TypeVar(hd listTypeVars); val alphaList = mkTypeConstruction ("list", listType, [alpha], [DeclaredAt inBasis]); val consType = mkFunctionType (mkProductType [alpha, alphaList], alphaList); val nilConstructor = makeValueConstr ("nil", alphaList, true, 2, Local{addr=ref ~1, level=ref baseLevel}, [DeclaredAt inBasis]) val consConstructor = makeValueConstr ("::", consType, false, 2, Local{addr=ref ~1, level=ref baseLevel}, [DeclaredAt inBasis]) val nilRepresentation = createNullaryConstructor(ShortForm 0w0, [alpha], "nil") val consRepresentation = createUnaryConstructor(UnboxedForm 2, [alpha], "::") in ([(nilConstructor, nilRepresentation), (consConstructor, consRepresentation)], mkInlproc(TypeValue.boxedEither, 1, "boxed-list", [], 0), mkInlproc(TypeValue.singleWord, 1, "size-list", [], 0)) end in val listConstr = buildBasisDatatype("list", "list", [makeTv {value=EmptyType, level=generalisable, nonunifiable=false, equality=false, printable=false}], true, makeConsAndNil) val (nilConstructor, consConstructor) = case listConstr of TypeConstrSet(_, [consC as Value{name="::", ...}, nilC as Value{name="nil", ...}]) => (nilC, consC) | _ => raise InternalError "nil and cons in wrong order" end local fun makeNoneAndSome optionType = let val optionTypeVars = tcTypeVars optionType; val alpha = TypeVar(hd optionTypeVars); val alphaOption = mkTypeConstruction ("option", optionType, [alpha], [DeclaredAt inBasis]); val someType = mkFunctionType (alpha, alphaOption); val noneConstructor = makeValueConstr ("NONE", alphaOption, true, 2, Local{addr=ref ~1, level=ref baseLevel}, [DeclaredAt inBasis]); val someConstructor = makeValueConstr ("SOME", someType, false, 2, Local{addr=ref ~1, level=ref baseLevel}, [DeclaredAt inBasis]); val noneRepresentation = createNullaryConstructor(ShortForm 0w0, [alpha], "NONE") and someRepresentation = createUnaryConstructor(BoxedForm, [alpha], "SOME") in ([(noneConstructor, noneRepresentation), (someConstructor, someRepresentation)], mkInlproc(TypeValue.boxedEither, 1, "boxed-option", [], 0), mkInlproc(TypeValue.singleWord, 1, "size-option", [], 0)) end in val optionConstr= buildBasisDatatype("option", "option", [makeTv {value=EmptyType, level=generalisable, nonunifiable=false, equality=false, printable=false}], true, makeNoneAndSome) val (noneConstructor, someConstructor) = case optionConstr of TypeConstrSet(_, [noneC as Value{name="NONE", ...}, someC as Value{name="SOME", ...}]) => (noneC, someC) | _ => raise InternalError "NONE and SOME in wrong order" end local fun listConstruct (base : types) : types = let val TypeConstrSet(listCons, _) = listConstr in mkTypeConstruction ("list", listCons, [base], [DeclaredAt inBasis]) end; val intTypeConstr = TYPESTRUCT.fixedIntType val stringTypeConstr = TYPESTRUCT.stringType val boolTypeConstr = TYPESTRUCT.boolType in local val fields = [ mkLabelEntry("file", stringTypeConstr), mkLabelEntry("startLine", intTypeConstr), mkLabelEntry("startPosition", intTypeConstr), mkLabelEntry("endLine", intTypeConstr), mkLabelEntry("endPosition", intTypeConstr) ] in val locationCons = makeTypeConstructor("location", [], makeTypeFunction(basisDescription "PolyML.location", ([], mkLabelled(sortLabels fields, true))), [DeclaredAt inBasis]) val locationConstr = TypeConstrSet(locationCons, []) end local (* Pretty print context information. *) fun makeConstructors typeconstr = let val contextType = mkTypeConstruction ("context", typeconstr, [], [DeclaredAt inBasis]) val locationType = mkTypeConstruction ("location", locationCons, [], [DeclaredAt inBasis]) val constrs = [ ("ContextLocation", locationType), ("ContextProperty", mkProductType[stringTypeConstr, stringTypeConstr])]; (* The representation of this datatype is given in Pretty.sml and must be the same as the representation that chooseConstrRepr will use. *) val numConstrs = List.length constrs fun makeCons(s,t) = makeValueConstr(s, mkFunctionType(t, contextType), false, numConstrs, Local{addr=ref ~1, level=ref baseLevel}, [DeclaredAt inBasis]) val {constrs=constrCode, boxed, size} = chooseConstrRepr(constrs, []) in (ListPair.zipEq(List.map makeCons constrs, constrCode), boxed, size) end in val contextConstr = buildBasisDatatype("context", "PolyML.context", [], false, makeConstructors) end local fun makeConstructors typeconstr = let val TypeConstrSet(contextCons, _) = contextConstr val prettyType = mkTypeConstruction ("pretty", typeconstr, [], [DeclaredAt inBasis]) val contextType = mkTypeConstruction ("context", contextCons, [], [DeclaredAt inBasis]) val constrs = [ ("PrettyBlock", mkProductType[intTypeConstr, boolTypeConstr, listConstruct contextType, listConstruct prettyType]), ("PrettyBreak", mkProductType[intTypeConstr, intTypeConstr]), ("PrettyLineBreak", EmptyType), ("PrettyString", stringTypeConstr), ("PrettyStringWithWidth", mkProductType[stringTypeConstr, intTypeConstr])]; (* The representation of this datatype is given in Pretty.sml and must be the same as the representation that chooseConstrRepr will use. *) val numConstrs = List.length constrs fun makeCons(s,t) = let val (ctype, nullary) = case t of EmptyType => (prettyType, true) | t => (mkFunctionType(t, prettyType), false) in makeValueConstr(s, ctype, nullary, numConstrs, Local{addr=ref ~1, level=ref baseLevel}, [DeclaredAt inBasis]) end val {constrs=constrCode, ...} = chooseConstrRepr(constrs, []) in (ListPair.zipEq(List.map makeCons constrs, constrCode), TypeValue.boxedEither, TypeValue.singleWord) end in val prettyConstr = buildBasisDatatype("pretty", "PolyML.pretty", [], false, makeConstructors) end end (* The representation of ptProperties is given in ExportTree.sml and must also match. *) (* Construct an exception identifier - This is a ref (so we can uniquely identify it) containing a print function for the type. *) local (* The exception identifier contains a value of type (exn*int->pretty) option. *) val TypeConstrSet(optionCons, _) = optionConstr and TypeConstrSet(prettyCons, _) = prettyConstr val exnPrinter = mkTypeConstruction ("option", optionCons, [ mkFunctionType( mkProductType[TYPESTRUCT.exnType, TYPESTRUCT.fixedIntType], mkTypeConstruction ("pretty", prettyCons, [], [DeclaredAt inBasis]) ) ], [DeclaredAt inBasis]) in fun mkExIden(ty, level, tvMap) = let (* Get the constructor tuple, select the constructor operation, apply it to the type. *) val makeSome = applyToInstance( [{value=exnPrinter, equality=false, printity=false}], level, tvMap, fn _ => mkInd(1, case someConstructor of Value { access, ...} => vaGlobal access)) val makeNone = applyToInstance( [{value=exnPrinter, equality=false, printity=false}], level, tvMap, fn _ => mkInd(1, case noneConstructor of Value { access, ...} => vaGlobal access)) val printerCode = case ty of FunctionType { arg, ...} => mkEval(makeSome, [printerForType(arg, level, tvMap)]) | _ => makeNone in refApplyCode printerCode end end (* Types that can be shared. *) structure Sharing = struct type codetree = codetree type types = types type values = values type typeConstrSet = typeConstrSet type typeId = typeId type typeVarForm = typeVarForm type typeVarMap = typeVarMap type level = level end end; diff --git a/mlsource/MLCompiler/INITIALISE_.ML b/mlsource/MLCompiler/INITIALISE_.ML index a2a1c530..15699664 100644 --- a/mlsource/MLCompiler/INITIALISE_.ML +++ b/mlsource/MLCompiler/INITIALISE_.ML @@ -1,1930 +1,1930 @@ (* Copyright (c) 2000 Cambridge University Technical Services Limited - Updated David C.J. Matthews 2008-9, 2012, 2013, 2015-19 + Updated David C.J. Matthews 2008-9, 2012, 2013, 2015-20 This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) (* Title: Initialise ML Global Declarations. Author: Dave Matthews,Cambridge University Computer Laboratory Copyright Cambridge University 1985 *) functor INITIALISE_ ( structure LEX: LEXSIG structure TYPETREE : TYPETREESIG structure STRUCTVALS : STRUCTVALSIG structure VALUEOPS : VALUEOPSSIG structure CODETREE : CODETREESIG structure EXPORTTREE: EXPORTTREESIG structure DATATYPEREP: DATATYPEREPSIG structure TYPEIDCODE: TYPEIDCODESIG structure MAKE: MAKESIG structure ADDRESS : AddressSig structure DEBUG: DEBUGSIG structure DEBUGGER : DEBUGGERSIG structure PRETTY : PRETTYSIG structure PRINTTABLE: PRINTTABLESIG structure MISC : sig val unescapeString : string -> string exception Conversion of string; (* string to int conversion failure *) end structure VERSION: sig val compilerVersion: string val versionNumber: int end structure UNIVERSALTABLE: sig type universal = Universal.universal type univTable type 'a tag = 'a Universal.tag val univLookup: univTable * 'a tag * string -> 'a option val fold: (string * universal * 'a -> 'a) -> 'a -> univTable -> 'a end sharing STRUCTVALS.Sharing = VALUEOPS.Sharing = TYPETREE.Sharing = EXPORTTREE.Sharing = PRETTY.Sharing = CODETREE.Sharing = MAKE.Sharing = ADDRESS = DATATYPEREP.Sharing = TYPEIDCODE.Sharing = DEBUGGER.Sharing = LEX.Sharing = PRINTTABLE.Sharing sharing STRUCTVALS.Sharing = VALUEOPS.Sharing = TYPETREE.Sharing = EXPORTTREE.Sharing = PRETTY.Sharing = CODETREE.Sharing = MAKE.Sharing = ADDRESS = DATATYPEREP.Sharing = TYPEIDCODE.Sharing = DEBUGGER.Sharing = LEX.Sharing = PRINTTABLE.Sharing = UNIVERSALTABLE ) : sig type gEnv val initGlobalEnv : {globalTable : gEnv, intIsArbitraryPrecision: bool } -> unit end = struct open STRUCTVALS; open TYPETREE open VALUEOPS; open CODETREE; open ADDRESS; open MAKE; open MISC; open EXPORTTREE open DATATYPEREP val intInfType = mkTypeConstruction ("int", intInfConstr, [], []) and realType = mkTypeConstruction ("real", realConstr, [], []) and charType = mkTypeConstruction ("char", charConstr, [], []) and wordType = mkTypeConstruction ("word", wordConstr, [], []) val declInBasis = [DeclaredAt inBasis] fun applyList _ [] = () | applyList f (h :: t) = (f h : unit; applyList f t); fun initGlobalEnv{globalTable : gEnv, intIsArbitraryPrecision: bool } = let val Env globalEnv = MAKE.gEnvAsEnv globalTable val enterGlobalValue = #enterVal globalEnv; val enterGlobalType = #enterType globalEnv; (* Some routines to help make the types. *) local (* careful - STRUCTVALS.intType differs from TYPETREE.intType *) open TYPETREE; in (* Make some type variables *) fun makeEqTV () = mkTypeVar (generalisable, true, false, false) fun makeTV () = mkTypeVar (generalisable, false, false, false) fun makePrintTV() = mkTypeVar (generalisable, false, false, true) fun makeTypeVariable() = makeTv {value=emptyType, level=generalisable, equality=false, nonunifiable=false, printable=false} (* Make some functions *) infixr 5 ->> fun a ->> b = mkFunctionType (a, b); infix 7 **; fun a ** b = mkProductType [a, b]; (* Type identifiers for the types of the declarations. *) val Int = if intIsArbitraryPrecision then intInfType else fixedIntType val String = stringType; val Bool = boolType; val Unit = unitType; val Char = charType; val Word = wordType; val Real = realType val Exn = exnType val mkTypeConstruction = mkTypeConstruction; val () = setPreferredInt(if intIsArbitraryPrecision then intInfConstr else fixedIntConstr) end; fun makePolymorphic(tvs, c) = let open TYPEIDCODE val tvs = List.filter(fn TypeVar tv => not justForEqualityTypes orelse tvEquality tv | _ => false) tvs in if null tvs then c else mkInlproc(c, List.length tvs, "", [], 0) end (* Function to make a type identifier with a pretty printer that just prints "?". None of the types are equality types so the equality function is empty. *) local fun monotypePrinter _ = PRETTY.PrettyString "?" in fun defaultEqAndPrintCode () = let open TypeValue val code = createTypeValue{ eqCode = CodeZero, printCode = mkConst (toMachineWord (ref monotypePrinter)), boxedCode = boxedEither (* Assume this for the moment *), sizeCode = singleWord } in Global (genCode(code, [], 0) ()) end end fun makeTypeAbbreviation(name, fullName, typeVars, typeResult, locations) = makeTypeConstructor( name, typeVars, makeTypeFunction(basisDescription fullName, (typeVars, typeResult)), locations) (* Make an opaque type and add it to an environment. *) fun makeAndDeclareOpaqueType(typeName, fullName, env) = let val typeconstr = makeTypeConstructor(typeName, [], makeFreeId(0, defaultEqAndPrintCode(), false, basisDescription fullName), declInBasis); in #enterType env (typeName, TypeConstrSet(typeconstr, [])); mkTypeConstruction (typeName, typeconstr, [], declInBasis) end; (* List of something *) fun List (base : types) : types = mkTypeConstruction ("list", tsConstr listConstr, [base], declInBasis); (* ref something *) fun Ref (base : types) : types = mkTypeConstruction ("ref", refConstr, [base], declInBasis); fun Option (base : types) : types = mkTypeConstruction ("option", tsConstr optionConstr, [base], declInBasis); (* Type-dependent functions. *) fun mkSpecialFun (name:string, typeof:types, opn: typeDependent) : values = makeOverloaded (name, typeof, opn); (* Overloaded functions. *) fun mkOverloaded (name:string) (typeof: types) : values = mkSpecialFun(name, typeof, TypeDep) (* Make a structure. Returns the table as an environment so that entries can be added to the structure. *) fun makeStructure(parentEnv, name) = let val str as Struct{signat=Signatures{tab, ...}, ...} = makeEmptyGlobal name val () = #enterStruct parentEnv (name, str) val Env env = makeEnv tab in env end val () = enterGlobalType ("unit", TypeConstrSet(unitConstr, [])); local val falseCons = mkGconstr ("false", Bool, createNullaryConstructor(EnumForm{tag=0w0, maxTag=0w1}, [], "false"), true, 2, declInBasis) val trueCons = mkGconstr ("true", Bool, createNullaryConstructor(EnumForm{tag=0w1, maxTag=0w1}, [], "true"), true, 2, declInBasis) val boolEnv = makeStructure(globalEnv, "Bool") (* Bool structure *) val notFn = mkGvar("not", Bool ->> Bool, mkUnaryFn BuiltIns.NotBoolean, declInBasis) in val () = #enterType boolEnv ("bool", TypeConstrSet(boolConstr, [trueCons, falseCons])) val () = #enterVal boolEnv ("true", trueCons) val () = #enterVal boolEnv ("false", falseCons) val () = #enterVal boolEnv ("not", notFn) end; val () = enterGlobalType ("int", TypeConstrSet(if intIsArbitraryPrecision then intInfConstr else fixedIntConstr, [])) val () = enterGlobalType ("char", TypeConstrSet(charConstr, [])) val () = enterGlobalType ("string", TypeConstrSet(stringConstr, [])) val () = enterGlobalType ("real", TypeConstrSet(realConstr, [])) val () = (* Enter :: and nil. *) List.app(fn(tv as Value{name, ...}) => enterGlobalValue(name, tv)) (tsConstructors listConstr) val () = enterGlobalType ("list", listConstr); val () = (* Enter NONE and SOME. *) List.app(fn(tv as Value{name, ...}) => enterGlobalValue(name, tv)) (tsConstructors optionConstr) val () = enterGlobalType ("option", optionConstr); local val refCons = let val a = mkTypeVar(generalisable, false, false, false) in mkGconstr ("ref", a ->> Ref a, createUnaryConstructor(RefForm, [a], "ref"), false, 1, declInBasis) end in val () = enterGlobalType ("ref", TypeConstrSet(refConstr, [refCons])); val () = enterGlobalValue ("ref", refCons); end local open BuiltIns fun monoTypePrinter _ = PRETTY.PrettyString "?" val idCode = let open TypeValue val equalLongWordFn = mkInlproc( mkBinary(LargeWordComparison TestEqual, mkLoadArgument 0, mkLoadArgument 1), 2, "EqualLargeWord()", [], 0) val code = createTypeValue{ eqCode=equalLongWordFn, printCode=mkConst (toMachineWord (ref monoTypePrinter)), boxedCode = boxedNever, sizeCode = singleWord } in Global (genCode(code, [], 0) ()) end in val largeWordType = makeTypeConstructor("word", [], makeFreeId(0, idCode, true, basisDescription "word"), declInBasis) val LargeWord = mkTypeConstruction ("LargeWord.word", largeWordType, [], declInBasis) end val () = enterGlobalType ("exn", TypeConstrSet(exnConstr, [])); val () = enterGlobalType ("word", TypeConstrSet(wordConstr, [])); val runCallEnv = makeStructure(globalEnv, "RunCall") fun enterRunCall (name : string, entry : codetree, typ : types) : unit = let val value = mkGvar (name, typ, entry, declInBasis); in #enterVal runCallEnv (name, value) end local (* unsafeCast. Can be used to convert any value to any type. *) val a = makeTV () val b = makeTV () val unsafeCastEntry = mkInlproc (mkLoadArgument 0 (* just the parameter *), 1, "unsafeCast(1)", [], 0) in val () = enterRunCall ("unsafeCast", makePolymorphic([a, b], unsafeCastEntry), a ->> b) end local val a = makeTV() and b = makeTV() open BuiltIns in (* isShort - test if a value is tagged rather than being an address. *) val () = enterRunCall ("isShort", makePolymorphic([a], mkUnaryFn IsTaggedValue), a ->> Bool) (* pointer equality *) val () = enterRunCall ("pointerEq", - makePolymorphic([a], mkBinaryFn(WordComparison{test=TestEqual, isSigned=false})), a ** a ->> Bool) + makePolymorphic([a], mkBinaryFn PointerEq), a ** a ->> Bool) (* load a word. The index is in words and is always zero or positive. *) val () = enterRunCall ("loadWord", makePolymorphic([a, b], mkLoadOperationFn(LoadStoreMLWord{isImmutable=false})), a ** Word ->> b) (* Load a word from an immutable. *) val () = enterRunCall ("loadWordFromImmutable", makePolymorphic([a, b], mkLoadOperationFn(LoadStoreMLWord{isImmutable=true})), a ** Word ->> b) (* load a byte. The index is in bytes and is always zero or positive. Probably the result should be a Word8.word value or a char. *) val () = enterRunCall ("loadByte", makePolymorphic([a, b], mkLoadOperationFn(LoadStoreMLByte{isImmutable=false})), a ** Word ->> b) (* Load a byte from an immutable. *) val () = enterRunCall ("loadByteFromImmutable", makePolymorphic([a, b], mkLoadOperationFn(LoadStoreMLByte{isImmutable=true})), a ** Word ->> b) (* Get the length of a heap cell. *) val () = enterRunCall ("memoryCellLength", makePolymorphic([a], mkUnaryFn MemoryCellLength), a ->> Word) (* Return the flags. Perhaps this could return a Word8.word value instead of a word. *) val () = enterRunCall ("memoryCellFlags", makePolymorphic([a], mkUnaryFn MemoryCellFlags), a ->> Word) (* Return the number of bytes per word. This is a constant since we have separate pre-built compilers for 32-bit and 64-bit. N.B. The byte order is not a constant since we only have a single pre-built compiler for little-endian and big-endian interpreted code. *) val () = enterRunCall ("bytesPerWord", mkConst(toMachineWord wordSize), Word) (* Store a word *) val () = enterRunCall ("storeWord", makePolymorphic([a, b], mkStoreOperationFn(LoadStoreMLWord{isImmutable=false})), mkProductType[a, Word, b] ->> Unit) (* Store a byte *) val () = enterRunCall ("storeByte", makePolymorphic([a, b], mkStoreOperationFn(LoadStoreMLByte{isImmutable=false})), mkProductType[a, Word, b] ->> Unit) (* Lock a mutable cell. *) val () = enterRunCall ("clearMutableBit", makePolymorphic([a], mkUnaryFn ClearMutableFlag), a ->> Unit) (* Allocate a byte cell. The second argument is the flags byte. It might be better if this were a Word8.word value. *) val () = enterRunCall ("allocateByteMemory", makePolymorphic([a], mkBinaryFn AllocateByteMemory), Word ** Word ->> a) (* Allocate a word cell. *) val () = enterRunCall ("allocateWordMemory", makePolymorphic([a, b], mkAllocateWordMemoryFn), mkProductType[Word, Word, a] ->> b) (* Byte vector operations. *) val () = enterRunCall ("byteVectorEqual", makePolymorphic([a], mkBlockOperationFn BlockOpEqualByte), mkProductType[a, a, Word, Word, Word] ->> Bool) val () = enterRunCall ("byteVectorCompare", makePolymorphic([a], mkBlockOperationFn BlockOpCompareByte), mkProductType[a, a, Word, Word, Word] ->> Int) (* Block moves. *) val () = enterRunCall ("moveBytes", makePolymorphic([a], mkBlockOperationFn (BlockOpMove{isByteMove=true})), mkProductType[a, a, Word, Word, Word] ->> Unit) val () = enterRunCall ("moveWords", makePolymorphic([a], mkBlockOperationFn (BlockOpMove{isByteMove=false})), mkProductType[a, a, Word, Word, Word] ->> Unit) (* Untagged loads and stores. *) val () = enterRunCall ("loadUntagged", mkLoadOperationFn LoadStoreUntaggedUnsigned, String ** Word ->> Word) val () = enterRunCall ("storeUntagged", mkStoreOperationFn LoadStoreUntaggedUnsigned, mkProductType[String, Word, Word] ->> Unit) val () = enterRunCall ("touch", makePolymorphic([a], mkUnaryFn TouchAddress), a ->> Unit) end local val debugOpts = [] (* Place to add debugging if necessary. *) (* [tagInject Pretty.compilerOutputTag (Pretty.prettyPrint(print, 70)), tagInject assemblyCodeTag true] *) fun makeCall rtsCall n entryName = rtsCall (entryName, n, debugOpts) val makeFastCall = makeCall CODETREE.Foreign.rtsCallFast and makeFullCall = makeCall CODETREE.Foreign.rtsCallFull (* We need to wrap this so that the arguments are passed in registers. *) fun makeRunCallTupled (argTypes, resultType, callN) : codetree = let val width = List.length argTypes val name = "rtsCall" ^ Int.toString width; local val f = mkLoadClosure 0 (* first item from enclosing scope *) val tuple = mkLoadArgument 0 (* the inner parameter *) val args = case argTypes of [singleType] => [(tuple, singleType)] | argTypes => let val argVals = List.tabulate(width, fn n => mkInd (n, tuple)) in ListPair.zipEq(argVals, argTypes) end in val innerBody = mkCall (f, args, resultType) end local (* The closure contains the address of the RTS call. *) val f = mkEval(mkConst callN, [mkLoadArgument 0]) val innerLambda = mkInlproc (innerBody, 1, name ^ "(1)", [mkLoadLocal 0], 0) in val outerBody = mkEnv([mkDec (0, f)], innerLambda) end val outerLambda = mkInlproc (outerBody, 1, name, [], 1) in outerLambda end local val a = makeTV () and b = makeTV () fun makeInlCode(makeCall, name) = let val call1 = toMachineWord(makeCall 1) val body = mkEval(mkConst call1, [mkLoadArgument 0]) val proc = mkInlproc (body, 1, name, [], 0) in makePolymorphic([a, b], proc) end in val rtsCallFull1Entry = makeInlCode(makeFullCall, "rtsCallFull1") and rtsCallFast1Entry = makeInlCode(makeFastCall, "rtsCallFast1") end local val a = makeTV () and b = makeTV () and c = makeTV () and d = makeTV () and e = makeTV () fun makeRtsCall(n, makeCall) = makeRunCallTupled(List.tabulate(n, fn _ => GeneralType), GeneralType, toMachineWord(makeCall n)) in val rtsCallFull0Entry = makePolymorphic([a], makeRtsCall(0, makeFullCall)) and rtsCallFast0Entry = makePolymorphic([a], makeRtsCall(0, makeFastCall)) val rtsCall0Type = String ->> Unit ->> a val rtsCall1Type = String ->> a ->> b val rtsCallFull2Entry = makePolymorphic([a, b, c], makeRtsCall(2, makeFullCall)) and rtsCallFast2Entry = makePolymorphic([a, b, c], makeRtsCall(2, makeFastCall)) val rtsCall2Type = String ->> TYPETREE.mkProductType [a,b] ->> c val rtsCallFull3Entry = makePolymorphic([a, b, c, d], makeRtsCall(3, makeFullCall)) val rtsCallFast3Entry = makePolymorphic([a, b, c, d], makeRtsCall(3, makeFastCall)) val rtsCall3Type = String ->> TYPETREE.mkProductType [a,b,c] ->> d val rtsCallFast4Entry = makePolymorphic([a, b, c, d, e], makeRtsCall(4, makeFastCall)) val rtsCall4Type = String ->> TYPETREE.mkProductType [a,b,c,d] ->> e end in val () = enterRunCall ("rtsCallFull0", rtsCallFull0Entry, rtsCall0Type) val () = enterRunCall ("rtsCallFast0", rtsCallFast0Entry, rtsCall0Type) val () = enterRunCall ("rtsCallFull1", rtsCallFull1Entry, rtsCall1Type) val () = enterRunCall ("rtsCallFast1", rtsCallFast1Entry, rtsCall1Type) val () = enterRunCall ("rtsCallFull2", rtsCallFull2Entry, rtsCall2Type) val () = enterRunCall ("rtsCallFast2", rtsCallFast2Entry, rtsCall2Type) val () = enterRunCall ("rtsCallFull3", rtsCallFull3Entry, rtsCall3Type) val () = enterRunCall ("rtsCallFast3", rtsCallFast3Entry, rtsCall3Type) val () = enterRunCall ("rtsCallFast4", rtsCallFast4Entry, rtsCall4Type) val makeRunCallTupled = makeRunCallTupled (* Needed for reals. *) end local (* Create nullary exception. *) fun makeException0(name, id) = let val exc = Value{ name = name, typeOf = TYPETREE.exnType, access = Global(mkConst(toMachineWord id)), class = Exception, locations = declInBasis, references = NONE, instanceTypes=NONE } in #enterVal runCallEnv (name, exc) end (* Create exception with parameter. *) and makeException1(name, id, exType) = let val exc = Value{ name = name, typeOf = exType ->> TYPETREE.exnType, access = Global(mkConst(toMachineWord id)), class = Exception, locations = declInBasis, references = NONE, instanceTypes=NONE } in #enterVal runCallEnv (name, exc) end (* Exception numbers. Most of these are hard-coded in the RTS. *) val EXC_interrupt = 1 val EXC_syserr = 2 val EXC_size = 4 val EXC_overflow = 5 val EXC_divide = 7 val EXC_conversion = 8 val EXC_XWindows = 10 val EXC_subscript = 11 val EXC_thread = 12 val EXC_foreign = 23 val EXC_Bind = 100 (* In Match compiler. *) val EXC_Match = 101 val EXC_Fail = 103 in val () = List.app makeException0 [ ("Interrupt", EXC_interrupt), ("Size", EXC_size), ("Bind", EXC_Bind), ("Div", EXC_divide), ("Match", EXC_Match), ("Overflow", EXC_overflow), ("Subscript", EXC_subscript) ] val () = List.app makeException1 [ ("Fail", EXC_Fail, String), ("Conversion", EXC_conversion, String), ("XWindows", EXC_XWindows, String), ("Foreign", EXC_foreign, String), ("Thread", EXC_thread, String), ("SysErr", EXC_syserr, String ** Option LargeWord) ] end (* Standard Basis structures for basic types. These contain the definitions of the basic types and operations on them. The structures are extended in the basis library and overloaded functions are extracted from them. *) local val largeIntEnv = makeStructure(globalEnv, "LargeInt") (* The comparison operations take two arbitrary precision ints and a general "compare" function that returns a fixed precision int. *) val compareType = mkProductType[intInfType, intInfType, intInfType ** intInfType ->> fixedIntType] ->> Bool val arithType = mkProductType[intInfType, intInfType, intInfType ** intInfType ->> intInfType] ->> intInfType fun enterArbitrary(name, oper, typ) = let val value = mkGvar (name, typ, mkArbitraryFn oper, declInBasis) in #enterVal largeIntEnv (name, value) end in val () = #enterType largeIntEnv ("int", TypeConstrSet(intInfConstr, [])) (* These functions are used internally. *) val () = enterArbitrary("less", ArbCompare BuiltIns.TestLess, compareType) val () = enterArbitrary("greater", ArbCompare BuiltIns.TestGreater, compareType) val () = enterArbitrary("lessEq", ArbCompare BuiltIns.TestLessEqual, compareType) val () = enterArbitrary("greaterEq", ArbCompare BuiltIns.TestGreaterEqual, compareType) val () = enterArbitrary("add", ArbArith BuiltIns.ArithAdd, arithType) val () = enterArbitrary("subtract", ArbArith BuiltIns.ArithSub, arithType) val () = enterArbitrary("multiply", ArbArith BuiltIns.ArithMult, arithType) end local val fixedIntEnv = makeStructure(globalEnv, "FixedInt") open BuiltIns fun enterBinary(name, oper, typ) = let val value = mkGvar (name, typ, mkBinaryFn oper, declInBasis) in #enterVal fixedIntEnv (name, value) end val compareType = fixedIntType ** fixedIntType ->> Bool and binaryType = fixedIntType ** fixedIntType ->> fixedIntType fun enterComparison(name, test) = enterBinary(name, WordComparison{test=test, isSigned=true}, compareType) and enterBinaryOp(name, oper) = enterBinary(name, FixedPrecisionArith oper, binaryType) in val () = #enterType fixedIntEnv ("int", TypeConstrSet(fixedIntConstr, [])) val () = enterComparison("<", TestLess) val () = enterComparison("<=", TestLessEqual) val () = enterComparison(">", TestGreater) val () = enterComparison(">=", TestGreaterEqual) val () = enterBinaryOp("+", ArithAdd) val () = enterBinaryOp("-", ArithSub) val () = enterBinaryOp("*", ArithMult) val () = enterBinaryOp("quot", ArithQuot) val () = enterBinaryOp("rem", ArithRem) end local open BuiltIns val largeWordEnv = makeStructure(globalEnv, "LargeWord") fun enterBinary(name, oper, typ) = let val value = mkGvar (name, typ, mkBinaryFn oper, declInBasis) in #enterVal largeWordEnv (name, value) end val compareType = LargeWord ** LargeWord ->> Bool and binaryType = LargeWord ** LargeWord ->> LargeWord and shiftType = LargeWord ** Word ->> LargeWord (* The shift amount is a Word. *) fun enterComparison(name, test) = enterBinary(name, LargeWordComparison test, compareType) and enterBinaryOp(name, oper) = enterBinary(name, LargeWordArith oper, binaryType) and enterBinaryLogical(name, oper) = enterBinary(name, LargeWordLogical oper, binaryType) and enterBinaryShift(name, oper) = enterBinary(name, LargeWordShift oper, shiftType) in val () = #enterType largeWordEnv ("word", TypeConstrSet(largeWordType, [])) val () = enterComparison("<", TestLess) val () = enterComparison("<=", TestLessEqual) val () = enterComparison(">", TestGreater) val () = enterComparison(">=", TestGreaterEqual) val () = enterBinaryOp("+", ArithAdd) val () = enterBinaryOp("-", ArithSub) val () = enterBinaryOp("*", ArithMult) val () = enterBinaryOp("div", ArithDiv) val () = enterBinaryOp("mod", ArithMod) val () = enterBinaryLogical("orb", LogicalOr) val () = enterBinaryLogical("andb", LogicalAnd) val () = enterBinaryLogical("xorb", LogicalXor) val () = enterBinaryShift("<<", ShiftLeft) val () = enterBinaryShift(">>", ShiftRightLogical) val () = enterBinaryShift("~>>", ShiftRightArithmetic) val LargeWord = LargeWord end local val wordStructEnv = makeStructure(globalEnv, "Word") open BuiltIns fun enterBinary(name, oper, typ) = let val value = mkGvar (name, typ, mkBinaryFn oper, declInBasis) in #enterVal wordStructEnv (name, value) end val compareType = Word ** Word ->> Bool and binaryType = Word ** Word ->> Word fun enterComparison(name, test) = enterBinary(name, WordComparison{test=test, isSigned=false}, compareType) and enterBinaryOp(name, oper) = enterBinary(name, WordArith oper, binaryType) and enterBinaryLogical(name, oper) = enterBinary(name, WordLogical oper, binaryType) and enterBinaryShift(name, oper) = enterBinary(name, WordShift oper, binaryType) val toLargeWordFn = mkGvar ("toLargeWord", Word ->> LargeWord, mkUnaryFn UnsignedToLongWord, declInBasis) and toLargeWordXFn = mkGvar ("toLargeWordX", Word ->> LargeWord, mkUnaryFn SignedToLongWord, declInBasis) and fromLargeWordFn = mkGvar ("fromLargeWord", LargeWord ->> Word, mkUnaryFn LongWordToTagged, declInBasis) in val () = #enterType wordStructEnv ("word", TypeConstrSet(wordConstr, [])) val () = enterComparison("<", TestLess) val () = enterComparison("<=", TestLessEqual) val () = enterComparison(">", TestGreater) val () = enterComparison(">=", TestGreaterEqual) val () = enterBinaryOp("+", ArithAdd) val () = enterBinaryOp("-", ArithSub) val () = enterBinaryOp("*", ArithMult) val () = enterBinaryOp("div", ArithDiv) val () = enterBinaryOp("mod", ArithMod) val () = enterBinaryLogical("orb", LogicalOr) val () = enterBinaryLogical("andb", LogicalAnd) val () = enterBinaryLogical("xorb", LogicalXor) val () = enterBinaryShift("<<", ShiftLeft) val () = enterBinaryShift(">>", ShiftRightLogical) val () = enterBinaryShift("~>>", ShiftRightArithmetic) val () = #enterVal wordStructEnv ("toLargeWord", toLargeWordFn) val () = #enterVal wordStructEnv ("toLargeWordX", toLargeWordXFn) val () = #enterVal wordStructEnv ("fromLargeWord", fromLargeWordFn) end local val charEnv = makeStructure(globalEnv, "Char") open BuiltIns (* Comparison functions are the same as Word. *) fun enterComparison(name, test) = let val typ = Char ** Char ->> Bool val entry = mkBinaryFn(WordComparison{test=test, isSigned=false}) val value = mkGvar (name, typ, entry, declInBasis) in #enterVal charEnv (name, value) end in val () = #enterType charEnv ("char", TypeConstrSet(charConstr, [])) val () = enterComparison("<", TestLess) val () = enterComparison("<=", TestLessEqual) val () = enterComparison(">", TestGreater) val () = enterComparison(">=", TestGreaterEqual) end local val stringEnv = makeStructure(globalEnv, "String") in val () = #enterType stringEnv ("string", TypeConstrSet(stringConstr, [])) end local val realEnv = makeStructure(globalEnv, "Real") (* These are only used in Real so are included here rather than in RunCall. rtsCallFastRealtoReal is used for functions such as sqrt. rtsCallFastGeneraltoReal is used for Real.fromLargeInt. *) val debugOpts = [] (* Place to add debugging if necessary. *) fun makeFastRealRealCall entryName = CODETREE.Foreign.rtsCallFastRealtoReal (entryName, debugOpts) and makeFastRealRealRealCall entryName = CODETREE.Foreign.rtsCallFastRealRealtoReal (entryName, debugOpts) and makeFastIntInfRealCall entryName = CODETREE.Foreign.rtsCallFastGeneraltoReal (entryName, debugOpts) and makeFastRealGeneralRealCall entryName = CODETREE.Foreign.rtsCallFastRealGeneraltoReal (entryName, debugOpts) val rtsCallFastR_REntry = makeRunCallTupled([DoubleFloatType], DoubleFloatType, toMachineWord makeFastRealRealCall) (* This needs to be tupled. *) val rtsCallFastRR_REntry = makeRunCallTupled([DoubleFloatType, DoubleFloatType], DoubleFloatType, toMachineWord makeFastRealRealRealCall) and rtsCallFastRI_REntry = makeRunCallTupled([DoubleFloatType, GeneralType], DoubleFloatType, toMachineWord makeFastRealGeneralRealCall) val rtsCallFastI_REntry = makeRunCallTupled([GeneralType], DoubleFloatType, toMachineWord makeFastIntInfRealCall) val rtsCallFastF_F = mkGvar ("rtsCallFastR_R", String ->> Real ->> Real, rtsCallFastR_REntry, declInBasis) val rtsCallFastFF_F = mkGvar ("rtsCallFastRR_R", String ->> Real ** Real ->> Real, rtsCallFastRR_REntry, declInBasis) val rtsCallFastFG_F = mkGvar ("rtsCallFastRI_R", String ->> Real ** Int ->> Real, rtsCallFastRI_REntry, declInBasis) val rtsCallFastG_F = mkGvar ("rtsCallFastI_R", String ->> intInfType ->> Real, rtsCallFastI_REntry, declInBasis) fun enterUnary(name, oper, typ) = let val value = mkGvar (name, typ, mkUnaryFn oper, declInBasis) in #enterVal realEnv (name, value) end fun enterBinary(name, oper, typ) = let val value = mkGvar (name, typ, mkBinaryFn oper, declInBasis) in #enterVal realEnv (name, value) end val compareType = Real ** Real ->> Bool and binaryType = Real ** Real ->> Real and unaryType = Real ->> Real and realToFixType = Real ->> fixedIntType open BuiltIns IEEEReal fun enterComparison(name, test) = enterBinary(name, RealComparison(test, PrecDouble), compareType) and enterBinaryOp(name, oper) = enterBinary(name, RealArith(oper, PrecDouble), binaryType) in val () = #enterType realEnv ("real", TypeConstrSet(realConstr, [])) val () = #enterVal realEnv ("rtsCallFastR_R", rtsCallFastF_F) val () = #enterVal realEnv ("rtsCallFastRR_R", rtsCallFastFF_F) val () = #enterVal realEnv ("rtsCallFastRI_R", rtsCallFastFG_F) val () = #enterVal realEnv ("rtsCallFastI_R", rtsCallFastG_F) val () = enterComparison("<", TestLess) val () = enterComparison("<=", TestLessEqual) val () = enterComparison(">", TestGreater) val () = enterComparison(">=", TestGreaterEqual) val () = enterComparison("==", TestEqual) (* real is not an eqtype. *) (* Included unordered mainly because it's easy to implement isNan. *) val () = enterComparison("unordered", TestUnordered) val () = enterBinaryOp("+", ArithAdd) val () = enterBinaryOp("-", ArithSub) val () = enterBinaryOp("*", ArithMult) val () = enterBinaryOp("/", ArithDiv) val () = enterUnary("~", RealNeg PrecDouble, unaryType) val () = enterUnary("abs", RealAbs PrecDouble, unaryType) val () = enterUnary("fromFixedInt", RealFixedInt PrecDouble, fixedIntType ->> Real) val () = enterUnary("truncFix", RealToInt(PrecDouble, TO_ZERO), realToFixType) val () = enterUnary("roundFix", RealToInt(PrecDouble, TO_NEAREST), realToFixType) val () = enterUnary("ceilFix", RealToInt(PrecDouble, TO_POSINF), realToFixType) val () = enterUnary("floorFix", RealToInt(PrecDouble, TO_NEGINF), realToFixType) end local val real32Env = makeStructure(globalEnv, "Real32") val floatType = mkTypeConstruction ("real", floatConstr, [], []) val Float = floatType val debugOpts = [] (* Place to add debugging if necessary. *) fun makeFastFloatFloatCall entryName = CODETREE.Foreign.rtsCallFastFloattoFloat (entryName, debugOpts) and makeFastFloatFloatFloatCall entryName = CODETREE.Foreign.rtsCallFastFloatFloattoFloat (entryName, debugOpts) and makeFastIntInfFloatCall entryName = CODETREE.Foreign.rtsCallFastGeneraltoFloat (entryName, debugOpts) and makeFastFloatGeneralFloatCall entryName = CODETREE.Foreign.rtsCallFastFloatGeneraltoFloat (entryName, debugOpts) val rtsCallFastR_REntry = makeRunCallTupled([SingleFloatType], SingleFloatType, toMachineWord makeFastFloatFloatCall) (* This needs to be tupled. *) val rtsCallFastRR_REntry = makeRunCallTupled([SingleFloatType, SingleFloatType], SingleFloatType, toMachineWord makeFastFloatFloatFloatCall) and rtsCallFastRI_REntry = makeRunCallTupled([SingleFloatType, GeneralType], SingleFloatType, toMachineWord makeFastFloatGeneralFloatCall) val rtsCallFastI_REntry = makeRunCallTupled([GeneralType], SingleFloatType, toMachineWord makeFastIntInfFloatCall) val rtsCallFastF_F = mkGvar ("rtsCallFastF_F", String ->> Float ->> Float, rtsCallFastR_REntry, declInBasis) val rtsCallFastFF_F = mkGvar ("rtsCallFastFF_F", String ->> Float ** Float ->> Float, rtsCallFastRR_REntry, declInBasis) val rtsCallFastFG_F = mkGvar ("rtsCallFastFI_F", String ->> Float ** Int ->> Float, rtsCallFastRI_REntry, declInBasis) val rtsCallFastG_F = mkGvar ("rtsCallFastI_F", String ->> intInfType ->> Float, rtsCallFastI_REntry, declInBasis) fun enterUnary(name, oper, typ) = let val value = mkGvar (name, typ, mkUnaryFn oper, declInBasis) in #enterVal real32Env (name, value) end fun enterBinary(name, oper, typ) = let val value = mkGvar (name, typ, mkBinaryFn oper, declInBasis) in #enterVal real32Env (name, value) end val compareType = Float ** Float ->> Bool and binaryType = Float ** Float ->> Float and unaryType = Float ->> Float and floatToFixType = Float ->> fixedIntType open BuiltIns IEEEReal fun enterComparison(name, test) = enterBinary(name, RealComparison(test, PrecSingle), compareType) and enterBinaryOp(name, oper) = enterBinary(name, RealArith(oper, PrecSingle), binaryType) in val () = #enterType real32Env ("real", TypeConstrSet(floatConstr, [])) val () = enterUnary("toLarge", BuiltIns.FloatToDouble, floatType ->> Real) (* Conversion with the current rounding mode. *) and () = enterUnary("fromReal", BuiltIns.DoubleToFloat NONE, Real ->> floatType) (* There are various versions of this function for each of the rounding modes. *) and () = enterUnary("fromRealRound", BuiltIns.DoubleToFloat (SOME IEEEReal.TO_NEAREST), Real ->> floatType) and () = enterUnary("fromRealTrunc", BuiltIns.DoubleToFloat (SOME IEEEReal.TO_ZERO), Real ->> floatType) and () = enterUnary("fromRealCeil", BuiltIns.DoubleToFloat (SOME IEEEReal.TO_POSINF), Real ->> floatType) and () = enterUnary("fromRealFloor", BuiltIns.DoubleToFloat (SOME IEEEReal.TO_NEGINF), Real ->> floatType) val () = #enterVal real32Env ("rtsCallFastR_R", rtsCallFastF_F) val () = #enterVal real32Env ("rtsCallFastRR_R", rtsCallFastFF_F) val () = #enterVal real32Env ("rtsCallFastRI_R", rtsCallFastFG_F) val () = #enterVal real32Env ("rtsCallFastI_R", rtsCallFastG_F) val () = enterComparison("<", TestLess) val () = enterComparison("<=", TestLessEqual) val () = enterComparison(">", TestGreater) val () = enterComparison(">=", TestGreaterEqual) val () = enterComparison("==", TestEqual) (* Real32.real is not an eqtype. *) val () = enterComparison("unordered", TestUnordered) val () = enterBinaryOp("+", ArithAdd) val () = enterBinaryOp("-", ArithSub) val () = enterBinaryOp("*", ArithMult) val () = enterBinaryOp("/", ArithDiv) val () = enterUnary("~", RealNeg PrecSingle, unaryType) val () = enterUnary("abs", RealAbs PrecSingle, unaryType) val () = enterUnary("truncFix", RealToInt(PrecSingle, TO_ZERO), floatToFixType) val () = enterUnary("roundFix", RealToInt(PrecSingle, TO_NEAREST), floatToFixType) val () = enterUnary("ceilFix", RealToInt(PrecSingle, TO_POSINF), floatToFixType) val () = enterUnary("floorFix", RealToInt(PrecSingle, TO_NEGINF), floatToFixType) end val bootstrapEnv = makeStructure(globalEnv, "Bootstrap") fun enterBootstrap (name : string, entry : codetree, typ : types) : unit = let val value = mkGvar (name, typ, entry, declInBasis) in #enterVal bootstrapEnv (name, value) end local val threadEnv = makeStructure(globalEnv, "Thread") open TypeValue fun monoTypePrinter _ = PRETTY.PrettyString "?" val code = createTypeValue{ - eqCode=equalWordFn, + eqCode=equalPointerOrWordFn, printCode=mkConst (toMachineWord (ref monoTypePrinter)), boxedCode=boxedAlways, sizeCode=singleWord } (* Thread.thread type. This is an equality type with pointer equality. *) val threadConstr= makeTypeConstructor ( "thread", [], makeFreeId(0, Global (genCode(code, [], 0) ()), true, basisDescription "thread"), [DeclaredAt inBasis]) val threadType = mkTypeConstruction ("thread", threadConstr, [], []); val selfFunction = mkGvar ("self", Unit ->> threadType, getCurrentThreadIdFn, declInBasis) val atIncrFunction = mkGvar("atomicIncr", Ref Word ->> Word, mkUnaryFn BuiltIns.AtomicIncrement, declInBasis) val atDecrFunction = mkGvar("atomicDecr", Ref Word ->> Word, mkUnaryFn BuiltIns.AtomicDecrement, declInBasis) val atResetFunction = mkGvar("atomicReset", Ref Word ->> Unit, mkUnaryFn BuiltIns.AtomicReset, declInBasis) in val () = #enterType threadEnv ("thread", TypeConstrSet(threadConstr, [])) val () = #enterVal threadEnv ("self", selfFunction) val () = #enterVal threadEnv ("atomicIncr", atIncrFunction) val () = #enterVal threadEnv ("atomicDecr", atDecrFunction) val () = #enterVal threadEnv ("atomicReset", atResetFunction) end local val fmemEnv = makeStructure(globalEnv, "ForeignMemory") val a = makeTV() (* We don't have Word8.word or Word32.word at this point so the easiest way to deal with this is to make them polymorphic. *) val get8Function = mkGvar("get8", LargeWord ** Word ->> a, makePolymorphic([a], mkLoadOperationFn LoadStoreC8), declInBasis) val get16Function = mkGvar("get16", LargeWord ** Word ->> Word, mkLoadOperationFn LoadStoreC16, declInBasis) val get32Function = mkGvar("get32", LargeWord ** Word ->> a, makePolymorphic([a], mkLoadOperationFn LoadStoreC32), declInBasis) val get64Function = mkGvar("get64", LargeWord ** Word ->> LargeWord, mkLoadOperationFn LoadStoreC64, declInBasis) val getFloatFunction = mkGvar("getFloat", LargeWord ** Word ->> Real, mkLoadOperationFn LoadStoreCFloat, declInBasis) val getDoubleFunction = mkGvar("getDouble", LargeWord ** Word ->> Real, mkLoadOperationFn LoadStoreCDouble, declInBasis) val set8Function = mkGvar("set8", mkProductType[LargeWord, Word, a] ->> Unit, makePolymorphic([a], mkStoreOperationFn LoadStoreC8), declInBasis) val set16Function = mkGvar("set16", mkProductType[LargeWord, Word, Word] ->> Unit, mkStoreOperationFn LoadStoreC16, declInBasis) val set32Function = mkGvar("set32", mkProductType[LargeWord, Word, a] ->> Unit, makePolymorphic([a], mkStoreOperationFn LoadStoreC32), declInBasis) val set64Function = mkGvar("set64", mkProductType[LargeWord, Word, LargeWord] ->> Unit, mkStoreOperationFn LoadStoreC64, declInBasis) val setFloatFunction = mkGvar("setFloat", mkProductType[LargeWord, Word, Real] ->> Unit, mkStoreOperationFn LoadStoreCFloat, declInBasis) val setDoubleFunction = mkGvar("setDouble", mkProductType[LargeWord, Word, Real] ->> Unit, mkStoreOperationFn LoadStoreCDouble, declInBasis) in val () = #enterVal fmemEnv ("get8", get8Function) val () = #enterVal fmemEnv ("get16", get16Function) val () = #enterVal fmemEnv ("get32", get32Function) val () = #enterVal fmemEnv ("get64", get64Function) val () = #enterVal fmemEnv ("getFloat", getFloatFunction) val () = #enterVal fmemEnv ("getDouble", getDoubleFunction) val () = #enterVal fmemEnv ("set8", set8Function) val () = #enterVal fmemEnv ("set16", set16Function) val () = #enterVal fmemEnv ("set32", set32Function) val () = #enterVal fmemEnv ("set64", set64Function) val () = #enterVal fmemEnv ("setFloat", setFloatFunction) val () = #enterVal fmemEnv ("setDouble", setDoubleFunction) end local fun addVal (name : string, value : 'a, typ : types) : unit = enterBootstrap (name, mkConst (toMachineWord value), typ) (* These are only used during the bootstrap phase. Replacements are installed once the appropriate modules of the basis library are compiled. *) fun intOfString s = let val radix = if String.size s >= 3 andalso String.substring(s, 0, 2) = "0x" orelse String.size s >= 4 andalso String.substring(s, 0, 3) = "~0x" then StringCvt.HEX else StringCvt.DEC in case StringCvt.scanString (Int.scan radix) s of NONE => raise Conversion "Invalid integer constant" | SOME res => res end fun wordOfString s = let val radix = if String.size s > 2 andalso String.sub(s, 2) = #"x" then StringCvt.HEX else StringCvt.DEC in case StringCvt.scanString (Word.scan radix) s of NONE => raise Conversion "Invalid word constant" | SOME res => res end open PRINTTABLE val convstringCode = genCode(mkConst(toMachineWord unescapeString), [], 0) () val convintCode = genCode(mkConst(toMachineWord intOfString), [], 0) () val convwordCode = genCode(mkConst(toMachineWord wordOfString), [], 0) () in (* Conversion overloads used to be set by the ML bootstrap code. It's simpler to do that here but to maintain compatibility with the 5.6 compiler we need to define these. Once we've rebuilt the compiler this can be removed along with the code that uses it. *) val () = addVal ("convStringName", "convString": string, String) val () = addVal ("convInt", intOfString : string -> int, String ->> intInfType) val () = addVal ("convWord", wordOfString : string -> word, String ->> Word) (* Convert a string, recognising and converting the escape codes. *) val () = addVal ("convString", unescapeString: string -> string, String ->> String) (* Flag to indicate which version of Int to compile *) val () = addVal ("intIsArbitraryPrecision", intIsArbitraryPrecision, Bool) (* Install the overloads now. *) val () = addOverload("convString", stringConstr, convstringCode) val () = addOverload("convInt", fixedIntConstr, convintCode) val () = addOverload("convInt", intInfConstr, convintCode) val () = addOverload("convWord", wordConstr, convwordCode) end (* The only reason we have vector here is to get equality right. We need vector to be an equality type and to have a specific equality function. *) local fun polyTypePrinter _ _ = PRETTY.PrettyString "?" (* The equality function takes the base equality type as an argument. The inner function takes two arguments which are the two vectors to compare, checks the lengths and if they're equal applies the base equality to each field. *) val eqCode = mkInlproc( mkProc( mkEnv([ (* Length of the items. *) mkDec(0, mkUnary(BuiltIns.MemoryCellLength, mkLoadArgument 0)), mkDec(1, mkUnary(BuiltIns.MemoryCellLength, mkLoadArgument 1)), mkMutualDecs[(2, (* Loop function. *) mkProc( mkIf( (* Finished? *) - mkEqualWord(mkLoadClosure 0, mkLoadArgument 0), + mkEqualTaggedWord(mkLoadClosure 0, mkLoadArgument 0), CodeTrue, (* Yes, all equal. *) mkIf( mkEval( TypeValue.extractEquality(mkLoadClosure 2), (* Base equality fn *) [ mkLoadOperation(LoadStoreMLWord{isImmutable=true}, mkLoadClosure 3, mkLoadArgument 0), mkLoadOperation(LoadStoreMLWord{isImmutable=true}, mkLoadClosure 4, mkLoadArgument 0) ]), mkEval(mkLoadClosure 1, (* Recursive call with index+1. *) [ mkBinary(BuiltIns.WordArith BuiltIns.ArithAdd, mkLoadArgument 0, mkConst(toMachineWord 1)) ]), CodeFalse (* Not equal elements - result false *) ) ), 1, "vector-loop", [mkLoadLocal 0 (* Length *), mkLoadLocal 2 (* Loop function *), mkLoadClosure 0 (* Base equality function *), mkLoadArgument 0 (* Vector 0 *), mkLoadArgument 1 (* Vector 1 *)], 0))] ], mkIf( (* Test the lengths. *) - mkEqualWord(mkLoadLocal 0, mkLoadLocal 1), + mkEqualTaggedWord(mkLoadLocal 0, mkLoadLocal 1), (* Equal - test the contents. *) mkEval(mkLoadLocal 2, [CodeZero]), CodeFalse (* Not same length- result false *) ) ), 2, "vector-eq", [mkLoadArgument 0], 3), 1, "vector-eq()", [], 0) val idCode = (* Polytype *) let open TypeValue val code = createTypeValue{ eqCode=eqCode, printCode=mkConst (toMachineWord (ref polyTypePrinter)), boxedCode=mkInlproc(boxedAlways, 1, "boxed-vector", [], 0), sizeCode=mkInlproc(singleWord, 1, "size-vector", [], 0)} in Global (genCode(code, [], 0) ()) end in val vectorType = makeTypeConstructor("vector", [makeTypeVariable()], makeFreeId(1, idCode, true, basisDescription "vector"), declInBasis) val () = enterGlobalType ("vector", TypeConstrSet(vectorType, [])) end (* We also need a type with byte-wise equality. *) local fun monoTypePrinter _ = PRETTY.PrettyString "?" (* This is a monotype equality function that takes two byte vectors and compares them byte-by-byte for equality. Because they are vectors of bytes it's unsafe to load the whole words which could look like addresses if the bottom bit happens to be zero. *) val eqCode = mkProc( mkEnv([ (* Length of the items. *) mkDec(0, mkUnary(BuiltIns.MemoryCellLength, mkLoadArgument 0)), mkDec(1, mkUnary(BuiltIns.MemoryCellLength, mkLoadArgument 1)) ], mkIf( (* Test the lengths. *) - mkEqualWord(mkLoadLocal 0, mkLoadLocal 1), + mkEqualTaggedWord(mkLoadLocal 0, mkLoadLocal 1), (* Equal - test the contents. *) mkEnv([ (* ByteVecEqual takes a byte length so we have to multiply by the number of bytes per word. *) mkDec(2, mkBinary(BuiltIns.WordArith BuiltIns.ArithMult, mkConst(toMachineWord RunCall.bytesPerWord), mkLoadLocal 0)) ], mkBlockOperation{kind=BlockOpEqualByte, leftBase=mkLoadArgument 0, rightBase=mkLoadArgument 1, leftIndex=CodeZero, rightIndex=CodeZero, length=mkLoadLocal 2}), CodeFalse (* Not same length- result false *) ) ), 2, "byteVector-eq", [], 3) val idCode = (* Polytype *) let open TypeValue val code = createTypeValue{ eqCode=eqCode, printCode=mkConst (toMachineWord (ref monoTypePrinter)), boxedCode=boxedAlways, sizeCode=singleWord} in Global (genCode(code, [], 0) ()) end in val byteVectorType = makeTypeConstructor("byteVector", [], makeFreeId(0, idCode, true, basisDescription "byteVector"), declInBasis) val () = #enterType bootstrapEnv ("byteVector", TypeConstrSet(byteVectorType, [])) end (* We also need array and Array2.array to be passed through here so that they have the special property of being eqtypes even if their argument is not. "array" is defined to be in the global environment. *) val () = enterGlobalType ("array", TypeConstrSet(arrayConstr, [])) val () = #enterType bootstrapEnv ("array", TypeConstrSet(array2Constr, [])) val () = #enterType bootstrapEnv ("byteArray", TypeConstrSet(byteArrayConstr, [])) (* "=', '<>', PolyML.print etc are type-specific function which appear to be polymorphic. The compiler recognises these and treats them specially. For (in)equality that means generating type-specific versions of the equality operations; for print etc that means printing in a type-specific way. They can become true polymorphic functions and lose their type-specificity. For (in)equality that means defaulting to structure equality which is normal and expected behaviour. For print etc that means losing the ability to print and just printing "?" so it's important to avoid that happening. "open" treats type-specific functions specially and retains the type-specificity. That's important to allow the prelude code to expand the PolyML structure. *) local val eqType = let val a = makeEqTV () in a ** a ->> Bool end val eqVal = mkSpecialFun("=", eqType, Equal) in val () = enterGlobalValue ("=", eqVal) end local val neqType = let val a = makeEqTV () in a ** a ->> Bool end val neqVal = mkSpecialFun("<>", neqType, NotEqual) in val () = enterGlobalValue ("<>", neqVal) end val polyMLEnv = makeStructure(globalEnv, "PolyML") val enterPolyMLVal = #enterVal polyMLEnv local (* This version of the environment must match that used in the NameSpace structure. *) open TYPETREE (* Create a new structure for them. *) val nameSpaceEnv = makeStructure(polyMLEnv, "NameSpace") (* Substructures. *) val valuesEnv = makeStructure(nameSpaceEnv, "Values") and typesEnv = makeStructure(nameSpaceEnv, "TypeConstrs") and fixesEnv = makeStructure(nameSpaceEnv, "Infixes") and structsEnv = makeStructure(nameSpaceEnv, "Structures") and sigsEnv = makeStructure(nameSpaceEnv, "Signatures") and functsEnv = makeStructure(nameSpaceEnv, "Functors") (* Types for the basic values. These are opaque. *) val valueVal = makeAndDeclareOpaqueType("value", "PolyML.NameSpace.Values.value", valuesEnv) (* Representation of the type of a value. *) val Types = makeAndDeclareOpaqueType("typeExpression", "PolyML.NameSpace.Values.typeExpression", valuesEnv) val typeVal = makeAndDeclareOpaqueType("typeConstr", "PolyML.NameSpace.TypeConstrs.typeConstr", typesEnv) val fixityVal = makeAndDeclareOpaqueType("fixity", "PolyML.NameSpace.Infixes.fixity", fixesEnv) val signatureVal = makeAndDeclareOpaqueType("signatureVal", "PolyML.NameSpace.Signatures.signatureVal", sigsEnv) val structureVal = makeAndDeclareOpaqueType("structureVal", "PolyML.NameSpace.Structures.structureVal", structsEnv) val functorVal = makeAndDeclareOpaqueType("functorVal", "PolyML.NameSpace.Functors.functorVal", functsEnv) (* nameSpace type. Labelled record. *) fun createFields(name, vType): { name: string, typeof: types} list = let val enterFun = String ** vType ->> Unit val lookupFun = String ->> Option vType val allFun = Unit ->> List (String ** vType) in [mkLabelEntry("enter" ^ name, enterFun), mkLabelEntry("lookup" ^ name, lookupFun), mkLabelEntry("all" ^ name, allFun)] end (* We have to use the same names as we use in the env type because we're passing "env" values through the bootstrap. *) val valTypes = [("Val", valueVal), ("Type", typeVal), ("Fix", fixityVal), ("Struct", structureVal), ("Sig", signatureVal), ("Funct", functorVal)] val fields = List.foldl (fn (p,l) => createFields p @ l) [] valTypes val recordType = makeTypeAbbreviation("nameSpace", "PolyML.NameSpace.nameSpace", [], mkLabelled(sortLabels fields, true), declInBasis); val () = #enterType nameSpaceEnv ("nameSpace", TypeConstrSet(recordType, [])); (* The result type of the compiler includes valueVal etc. *) val resultFields = List.map TYPETREE.mkLabelEntry [("values", List(String ** valueVal)), ("fixes", List(String ** fixityVal)), ("types", List(String ** typeVal)), ("structures", List(String ** structureVal)), ("signatures", List(String ** signatureVal)), ("functors", List(String ** functorVal))] in val nameSpaceType = mkTypeConstruction ("nameSpace", recordType, [], declInBasis) val execResult = mkLabelled(sortLabels resultFields, true) type execResult = { fixes: (string * fixStatus) list, values: (string * values) list, structures: (string * structVals) list, signatures: (string * signatures) list, functors: (string * functors) list, types: (string * typeConstrSet) list } val valueVal = valueVal val typeVal = typeVal val fixityVal = fixityVal val signatureVal = signatureVal val structureVal = structureVal val functorVal = functorVal val Types = Types val valuesEnv = valuesEnv and typesEnv = typesEnv and fixesEnv = fixesEnv and structsEnv = structsEnv and sigsEnv = sigsEnv and functsEnv = functsEnv end local val typeconstr = locationConstr val () = #enterType polyMLEnv ("location", typeconstr); in val Location = mkTypeConstruction ("location", tsConstr typeconstr, [], declInBasis) end (* Interface to the debugger. *) local open TYPETREE val debuggerEnv = makeStructure(polyMLEnv, "DebuggerInterface") (* Make these opaque at this level. *) val locationPropList = makeAndDeclareOpaqueType("locationPropList", "PolyML.DebuggerInterface.locationPropList", debuggerEnv) val typeId = makeAndDeclareOpaqueType("typeId", "PolyML.DebuggerInterface.typeId", debuggerEnv) val machineWordType = makeAndDeclareOpaqueType("machineWord", "PolyML.DebuggerInterface.machineWord", debuggerEnv) (* For long term security keep these as different from global types and sigs. Values in the static environment need to be copied before they are global. *) val localType = makeAndDeclareOpaqueType("localType", "PolyML.DebuggerInterface.localType", debuggerEnv) val localTypeConstr = makeAndDeclareOpaqueType("localTypeConstr", "PolyML.DebuggerInterface.localTypeConstr", debuggerEnv) val localSig = makeAndDeclareOpaqueType("localSig", "PolyML.DebuggerInterface.localSig", debuggerEnv) open DEBUGGER (* Entries in the static list. This type is only used within the implementation of DebuggerInterface in the basis library and does not appear in the final signature. *) val environEntryConstr = makeTypeConstructor("environEntry", [], makeFreeId(0, defaultEqAndPrintCode(), false, basisDescription "PolyML.DebuggerInterface.environEntry"), declInBasis) val environEntryType = mkTypeConstruction ("environEntry", environEntryConstr, [], declInBasis) val constrs = (* Order is significant. *) [ ("EnvEndFunction", mkProductType[String, Location, localType]), ("EnvException", mkProductType[String, localType, locationPropList]), ("EnvStartFunction", mkProductType[String, Location, localType]), ("EnvStructure", mkProductType[String, localSig, locationPropList]), ("EnvTConstr", String ** localTypeConstr), ("EnvTypeid", typeId ** typeId), ("EnvVConstr", mkProductType[String, localType, Bool, Int, locationPropList]), ("EnvValue", mkProductType[String, localType, locationPropList]) ] (* This representation must match the representation defined in DEBUGGER_.sml. *) val numConstrs = List.length constrs val {constrs=constrReps, ...} = chooseConstrRepr(constrs, []) val constructors = ListPair.map (fn ((s,t), code) => mkGconstr(s, t ->> environEntryType, code, false, numConstrs, declInBasis)) (constrs, constrReps) val () = List.app (fn c => #enterVal debuggerEnv(valName c, c)) constructors (* Put these constructors onto the type. *) val () = #enterType debuggerEnv ("environEntry", TypeConstrSet(environEntryConstr, constructors)) (* Debug state type. *) val debugStateConstr = makeTypeAbbreviation("debugState", "PolyML.DebuggerInterface.debugState", [], mkProductType[List environEntryType, List machineWordType, Location], declInBasis) val () = #enterType debuggerEnv ("debugState", TypeConstrSet(debugStateConstr, [])) val debugStateType = mkTypeConstruction ("debugState", debugStateConstr, [], declInBasis) in val () = applyList (fn (name, v, t) => #enterVal debuggerEnv (name, mkGvar (name, t, mkConst v, declInBasis))) [ ("makeValue", toMachineWord(makeValue: debugState -> string * types * locationProp list * machineWord -> values), debugStateType ->> mkProductType[String, localType, locationPropList, machineWordType] ->> valueVal), ("makeException", toMachineWord(makeException: debugState -> string * types * locationProp list * machineWord -> values), debugStateType ->> mkProductType[String, localType, locationPropList, machineWordType] ->> valueVal), ("makeConstructor", toMachineWord(makeConstructor: debugState -> string * types * bool * int * locationProp list * machineWord -> values), debugStateType ->> mkProductType[String, localType, Bool, Int, locationPropList, machineWordType] ->> valueVal), ("makeAnonymousValue", toMachineWord(makeAnonymousValue: debugState -> types * machineWord -> values), debugStateType ->> mkProductType[localType, machineWordType] ->> valueVal), ("makeStructure", toMachineWord(makeStructure: debugState -> string * signatures * locationProp list * machineWord -> structVals), debugStateType ->> mkProductType[String, localSig, locationPropList, machineWordType] ->> structureVal), ("makeTypeConstr", toMachineWord(makeTypeConstr: debugState -> typeConstrSet -> typeConstrSet), debugStateType ->> localTypeConstr ->> typeVal), ("unitValue", toMachineWord(mkGvar("", unitType, CodeZero, []): values), valueVal), (* Used as a default *) ("setOnEntry", toMachineWord(setOnEntry: (string * PolyML.location -> unit) option -> unit), Option (String ** Location ->> Unit) ->> Unit), ("setOnExit", toMachineWord(setOnExit: (string * PolyML.location -> unit) option -> unit), Option (String ** Location ->> Unit) ->> Unit), ("setOnExitException", toMachineWord(setOnExitException: (string * PolyML.location -> exn -> unit) option -> unit), Option (String ** Location ->> Exn ->> Unit) ->> Unit), ("setOnBreakPoint", toMachineWord(setOnBreakPoint: (PolyML.location * bool ref -> unit) option -> unit), Option (Location ** Ref Bool ->> Unit) ->> Unit) ] end local val typeconstr = contextConstr in val () = #enterType polyMLEnv ("context", typeconstr); val () = List.app(fn(tv as Value{name, ...}) => #enterVal polyMLEnv(name, tv)) (tsConstructors typeconstr) end local val typeconstr = prettyConstr in val () = #enterType polyMLEnv ("pretty", typeconstr); val () = List.app(fn(tv as Value{name, ...}) => #enterVal polyMLEnv(name, tv)) (tsConstructors typeconstr) val PrettyType = mkTypeConstruction ("pretty", tsConstr typeconstr, [], declInBasis) end local val printType = let val a = makePrintTV () in a ->> a end; val printVal = mkSpecialFun("print", printType, Print); in val () = enterPolyMLVal ("print", printVal); end; local val makeStringType = let val a = makePrintTV () in a ->> String end; val makeStringVal = mkSpecialFun("makestring", makeStringType, MakeString); in val () = enterPolyMLVal ("makestring", makeStringVal); end; local val prettyType = let val a = makePrintTV () in a ** fixedIntType ->> PrettyType end; val prettyVal = mkSpecialFun("prettyRepresentation", prettyType, GetPretty); in val () = enterPolyMLVal ("prettyRepresentation", prettyVal); end; local (* addPrettyPrinter is the new function to install a pretty printer. *) val a = makeTV () val b = makeTV () val addPrettyType = (TYPETREE.fixedIntType ->> b ->> a ->> PrettyType) ->> Unit; val addPrettyVal = mkSpecialFun("addPrettyPrinter", addPrettyType, AddPretty); in val () = enterPolyMLVal ("addPrettyPrinter", addPrettyVal); end; (* This goes in RunCall since it's only for the basis library. *) local val addOverloadType = let val a = makeTV () and b = makeTV () in (a ->> b) ->> String ->> Unit end; val addOverloadVal = mkSpecialFun("addOverload", addOverloadType, AddOverload); in val () = #enterVal runCallEnv ("addOverload", addOverloadVal); end local (* Add a function to switch the default integer type. *) fun setType isArbitrary = setPreferredInt(if isArbitrary then intInfConstr else fixedIntConstr) in val () = #enterVal runCallEnv ("setDefaultIntTypeArbitrary", mkGvar ("setDefaultIntTypeArbitrary", Bool ->> Unit, mkConst (toMachineWord setType), declInBasis)) end local val sourceLocVal = mkSpecialFun("sourceLocation", Unit ->> Location, GetLocation); in val () = enterPolyMLVal ("sourceLocation", sourceLocVal); end; local (* This is used as one of the arguments to the compiler function. *) open TYPETREE val uniStructEnv = makeStructure(bootstrapEnv, "Universal") fun enterUniversal (name : string, entry : codetree, typ : types) : unit = let val value = mkGvar (name, typ, entry, declInBasis); in #enterVal uniStructEnv (name, value) end; local fun polyTypePrinter _ _ = PRETTY.PrettyString "?" open TypeValue val idCode = let val code = createTypeValue{ eqCode=CodeZero, (* Not an equality type *) printCode=mkConst (toMachineWord (ref polyTypePrinter)), boxedCode=mkInlproc(boxedEither(* Assume worst case *), 1, "boxed-tag", [], 0), sizeCode=mkInlproc(singleWord, 1, "size-tag", [], 0)} in Global (genCode(code, [], 0) ()) end in (* type 'a tag *) val tagConstr = makeTypeConstructor("tag", [makeTypeVariable()], makeFreeId(1, idCode, false, basisDescription "tag"), declInBasis); val () = #enterType uniStructEnv ("tag", TypeConstrSet(tagConstr, [])) end (* type universal *) val univConstr = makeTypeConstructor("universal", [], makeFreeId(0, defaultEqAndPrintCode(), false, basisDescription "universal"), declInBasis); val () = #enterType uniStructEnv ("universal", TypeConstrSet(univConstr, [])); fun Tag base = mkTypeConstruction ("tag", tagConstr, [base], declInBasis) val Universal = mkTypeConstruction ("universal", univConstr, [], declInBasis) val a = makeTV() (* val tagInject : 'a tag -> 'a -> universal *) val injectType = Tag a ->> a ->> Universal val () = enterUniversal ("tagInject", makePolymorphic([a], mkConst (toMachineWord (Universal.tagInject: 'a Universal.tag -> 'a -> Universal.universal))), injectType) (* We don't actually need tagIs and tagProject since this is only used for the compiler. Universal is redefined in the basis library. *) val projectType = Tag a ->> Universal ->> a val () = enterUniversal ("tagProject", makePolymorphic([a], mkConst (toMachineWord(Universal.tagProject: 'a Universal.tag -> Universal.universal -> 'a))), projectType) val testType = Tag a ->> Universal ->> Bool val () = enterUniversal ("tagIs", makePolymorphic([a], mkConst (toMachineWord(Universal.tagIs: 'a Universal.tag -> Universal.universal -> bool))), testType) in val Tag = Tag and Universal = Universal end local open TYPETREE (* Parsetree properties datatype. *) val propConstr = makeTypeConstructor("ptProperties", [], makeFreeId(0, defaultEqAndPrintCode(), false, basisDescription "PolyML.ptProperties"), declInBasis); val PtProperties = mkTypeConstruction ("ptProperties", propConstr, [], declInBasis) (* Parsetree type. *) val parseTreeConstr = makeTypeAbbreviation("parseTree", "PolyML.parseTree", [], Location ** List PtProperties, declInBasis); val ParseTree = mkTypeConstruction ("parseTree", parseTreeConstr, [], declInBasis) val () = #enterType polyMLEnv ("parseTree", TypeConstrSet(parseTreeConstr, [])); val constrs = (* Order is significant. *) [ ("PTbreakPoint", Ref Bool), ("PTcompletions", List String), ("PTdeclaredAt", Location), ("PTdefId", fixedIntType), ("PTfirstChild", Unit ->> ParseTree), ("PTnextSibling", Unit ->> ParseTree), ("PTopenedAt", Location), ("PTparent", Unit ->> ParseTree), ("PTpreviousSibling", Unit ->> ParseTree), ("PTprint", fixedIntType ->> PrettyType), ("PTreferences", Bool ** List Location), ("PTrefId", fixedIntType), ("PTstructureAt", Location), ("PTtype", Types) ]; (* This representation must match the representation defined in ExportTree.sml. *) val numConstrs = List.length constrs val {constrs=constrReps, ...} = chooseConstrRepr(constrs, []) val constructors = ListPair.map (fn ((s,t), code) => mkGconstr(s, t ->> PtProperties, code, false, numConstrs, declInBasis)) (constrs, constrReps) val () = List.app (fn c => #enterVal polyMLEnv(valName c, c)) constructors (* Put these constructors onto the type. *) val () = #enterType polyMLEnv ("ptProperties", TypeConstrSet(propConstr, constructors)); in val ParseTree = ParseTree and PtProperties = PtProperties end local open TYPETREE val compilerType : types = mkProductType[nameSpaceType, Unit ->> Option Char, List Universal] ->> mkProductType[Option ParseTree, Option (Unit ->> execResult)] type compilerType = nameSpace * (unit -> char option) * Universal.universal list -> exportTree option * (unit->execResult) option in val () = enterBootstrap ("use", mkConst (toMachineWord ((useIntoEnv globalTable []): string -> unit)), String ->> Unit) val () = enterBootstrap ("useWithParms", mkConst (toMachineWord ((useIntoEnv globalTable): Universal.universal list -> string -> unit)), List Universal ->> String ->> Unit) val () = enterPolyMLVal("compiler", mkGvar ("compiler", compilerType, mkConst (toMachineWord (compiler: compilerType)), declInBasis)); val () = enterBootstrap("globalSpace", mkConst (toMachineWord(gEnvAsNameSpace globalTable: nameSpace)), nameSpaceType) end; local val ty = TYPETREE.mkOverloadSet[] val addType = ty ** ty ->> ty; val negType = ty ->> ty; val cmpType = ty ** ty ->> Bool; in val () = enterGlobalValue ("+", mkOverloaded "+" addType); val () = enterGlobalValue ("-", mkOverloaded "-" addType); val () = enterGlobalValue ("*", mkOverloaded "*" addType); val () = enterGlobalValue ("~", mkOverloaded "~" negType); val () = enterGlobalValue ("abs", mkOverloaded "abs" negType); val () = enterGlobalValue (">=", mkOverloaded ">=" cmpType); val () = enterGlobalValue ("<=", mkOverloaded "<=" cmpType); val () = enterGlobalValue (">", mkOverloaded ">" cmpType); val () = enterGlobalValue ("<", mkOverloaded "<" cmpType); (* The following overloads are added in ML97 *) val () = enterGlobalValue ("div", mkOverloaded "div" addType); val () = enterGlobalValue ("mod", mkOverloaded "mod" addType); val () = enterGlobalValue ("/", mkOverloaded "/" addType); end; local open DEBUG; local open TYPETREE val fields = [ mkLabelEntry("location", Location), mkLabelEntry("hard", Bool), mkLabelEntry("message", PrettyType), mkLabelEntry("context", Option PrettyType) ] in val errorMessageProcType = mkLabelled(sortLabels fields, true) ->> Unit type errorMessageProcType = { location: location, hard: bool, message: pretty, context: pretty option } -> unit end local open TYPETREE val optNav = Option(Unit->>ParseTree) val fields = [ mkLabelEntry("parent", optNav), mkLabelEntry("next", optNav), mkLabelEntry("previous", optNav) ] in val navigationType = mkLabelled(sortLabels fields, true) type navigationType = { parent: (unit->exportTree) option, next: (unit->exportTree) option, previous: (unit->exportTree) option } end type 'a tag = 'a Universal.tag in val () = applyList (fn (name, v, t) => enterBootstrap(name, mkConst v, t)) [ ("compilerVersion", toMachineWord (VERSION.compilerVersion: string), String), ("compilerVersionNumber", toMachineWord (VERSION.versionNumber: int), Int), ("lineNumberTag", toMachineWord (lineNumberTag : (unit->FixedInt.int) tag), Tag (Unit->>fixedIntType)), ("offsetTag", toMachineWord (offsetTag: (unit->FixedInt.int) tag), Tag (Unit->>fixedIntType)), ("fileNameTag", toMachineWord (fileNameTag: string tag), Tag String), ("bindingCounterTag", toMachineWord (bindingCounterTag: (unit->FixedInt.int) tag), Tag (Unit->>fixedIntType)), ("maxInlineSizeTag", toMachineWord (maxInlineSizeTag: FixedInt.int tag), Tag fixedIntType), ("assemblyCodeTag", toMachineWord (assemblyCodeTag: bool tag), Tag Bool), ("parsetreeTag", toMachineWord (parsetreeTag: bool tag), Tag Bool), ("codetreeTag", toMachineWord (codetreeTag: bool tag), Tag Bool), ("icodeTag", toMachineWord (icodeTag: bool tag), Tag Bool), ("lowlevelOptimiseTag", toMachineWord (lowlevelOptimiseTag: bool tag), Tag Bool), ("codetreeAfterOptTag", toMachineWord (codetreeAfterOptTag: bool tag), Tag Bool), ("inlineFunctorsTag", toMachineWord (inlineFunctorsTag: bool tag), Tag Bool), ("debugTag", toMachineWord (debugTag: bool tag), Tag Bool), ("printDepthFunTag", toMachineWord (DEBUG.printDepthFunTag: (unit->FixedInt.int) tag), Tag (Unit->>fixedIntType)), ("errorDepthTag", toMachineWord (DEBUG.errorDepthTag: FixedInt.int tag), Tag fixedIntType), ("lineLengthTag", toMachineWord (DEBUG.lineLengthTag: FixedInt.int tag), Tag fixedIntType), ("profileAllocationTag", toMachineWord (DEBUG.profileAllocationTag: FixedInt.int tag), Tag fixedIntType), ("printOutputTag", toMachineWord (PRETTY.printOutputTag: (pretty->unit) tag), Tag (PrettyType->>Unit)) , ("compilerOutputTag", toMachineWord (PRETTY.compilerOutputTag: (pretty->unit) tag), Tag (PrettyType->>Unit)), ("errorMessageProcTag", toMachineWord (LEX.errorMessageProcTag: errorMessageProcType tag), Tag errorMessageProcType), ("rootTreeTag", toMachineWord (EXPORTTREE.rootTreeTag: navigation tag), Tag navigationType), ("reportUnreferencedIdsTag", toMachineWord (reportUnreferencedIdsTag: bool tag), Tag Bool), ("reportExhaustiveHandlersTag", toMachineWord (reportExhaustiveHandlersTag: bool tag), Tag Bool), ("narrowOverloadFlexRecordTag", toMachineWord (narrowOverloadFlexRecordTag: bool tag), Tag Bool), ("createPrintFunctionsTag", toMachineWord (createPrintFunctionsTag: bool tag), Tag Bool), ("reportDiscardedValuesTag", toMachineWord (reportDiscardedValuesTag: FixedInt.int tag), Tag fixedIntType) ] end; (* PolyML.CodeTree structure. This exports the CodeTree structure into the ML space. *) local open CODETREE val codetreeEnv = makeStructure(polyMLEnv, "CodeTree") fun createType typeName = makeAndDeclareOpaqueType(typeName, "PolyML.CodeTree." ^ typeName, codetreeEnv) val CodeTree = createType "codetree" and MachineWord = createType "machineWord" and CodeBinding = createType "codeBinding" (* For the moment export these only for the general argument and result types. *) fun simpleFn (code, nArgs, name, closure, nLocals) = mkFunction{body=code, argTypes=List.tabulate(nArgs, fn _ => GeneralType), resultType=GeneralType, name=name, closure=closure, numLocals=nLocals} and simpleInlineFn (code, nArgs, name, closure, nLocals) = mkInlineFunction{body=code, argTypes=List.tabulate(nArgs, fn _ => GeneralType), resultType=GeneralType, name=name, closure=closure, numLocals=nLocals} and simpleCall(func, args) = mkCall(func, List.map (fn c => (c, GeneralType)) args, GeneralType) in val CodeTree = CodeTree val () = applyList (fn (name, v, t) => #enterVal codetreeEnv (name, mkGvar (name, t, mkConst v, declInBasis))) [ ("pretty", toMachineWord (CODETREE.pretty: codetree -> pretty), CodeTree ->> PrettyType), ("mkConstant", toMachineWord(mkConst: machineWord -> codetree), MachineWord ->> CodeTree), ("genCode", toMachineWord (genCode: codetree * Universal.universal list * int -> (unit->codetree)), mkProductType[CodeTree, List Universal, Int] ->> (Unit ->> CodeTree)), ("evalue", toMachineWord (evalue: codetree -> machineWord option), CodeTree ->> Option MachineWord), ("mkFunction", toMachineWord (simpleFn: codetree * int * string * codetree list * int -> codetree), mkProductType[CodeTree, Int, String, List CodeTree, Int] ->> CodeTree), ("mkInlineFunction", toMachineWord (simpleInlineFn: codetree * int * string * codetree list * int -> codetree), mkProductType[CodeTree, Int, String, List CodeTree, Int] ->> CodeTree), ("mkCall", toMachineWord (simpleCall: codetree * codetree list -> codetree), CodeTree ** List CodeTree ->> CodeTree), ("mkLoadLocal", toMachineWord (mkLoadLocal: int -> codetree), Int ->> CodeTree), ("mkLoadArgument", toMachineWord (mkLoadArgument: int -> codetree), Int ->> CodeTree), ("mkLoadClosure", toMachineWord (mkLoadClosure: int -> codetree), Int ->> CodeTree), ("mkDec", toMachineWord (mkDec: int * codetree -> codeBinding), Int ** CodeTree ->> CodeBinding), ("mkInd", toMachineWord (mkInd: int * codetree -> codetree), Int ** CodeTree ->> CodeTree), ("mkIf", toMachineWord (mkIf: codetree * codetree * codetree -> codetree), mkProductType[CodeTree, CodeTree, CodeTree] ->> CodeTree), ("mkWhile", toMachineWord (mkWhile: codetree * codetree -> codetree), CodeTree ** CodeTree ->> CodeTree), ("mkLoop", toMachineWord (mkLoop: codetree list -> codetree), List CodeTree ->> CodeTree), ("mkBeginLoop", toMachineWord (mkBeginLoop: codetree * (int * codetree) list -> codetree), CodeTree ** List(Int ** CodeTree) ->> CodeTree), ("mkEnv", toMachineWord (mkEnv: codeBinding list * codetree -> codetree), List CodeBinding ** CodeTree ->> CodeTree), ("mkMutualDecs", toMachineWord (mkMutualDecs: (int * codetree) list -> codeBinding), List(Int ** CodeTree) ->> CodeBinding), ("mkTuple", toMachineWord (mkTuple: codetree list -> codetree), List CodeTree ->> CodeTree), ("mkRaise", toMachineWord (mkRaise: codetree -> codetree), CodeTree ->> CodeTree), ("mkHandle", toMachineWord (mkHandle: codetree * codetree * int -> codetree), mkProductType[CodeTree, CodeTree, Int] ->> CodeTree), ("mkNullDec", toMachineWord (mkNullDec: codetree -> codeBinding), CodeTree ->> CodeBinding) ] end local (* Finish off the NameSpace structure now we have types such as pretty. *) open TYPETREE (* The exported versions expect full name spaces as arguments. Because we convert the exported versions to machineWord and give them types as data structures the compiler can't actually check that the type we give matched the internal type. *) fun makeTypeEnv NONE = { lookupType = fn _ => NONE, lookupStruct = fn _ => NONE } | makeTypeEnv(SOME(nameSpace: nameSpace)): printTypeEnv = { lookupType = fn s => case #lookupType nameSpace s of NONE => NONE | SOME t => SOME(t, NONE), lookupStruct = fn s => case #lookupStruct nameSpace s of NONE => NONE | SOME t => SOME(t, NONE) } local (* Values substructure. This also has operations related to type expressions. *) fun codeForValue (Value{access = Global code, class = ValBound, ...}) = code | codeForValue _ = raise Fail "Not a global value" and exportedDisplayTypeExp(ty, depth, nameSpace: nameSpace option) = TYPETREE.display(ty, depth, makeTypeEnv nameSpace) and exportedDisplayValues(valu, depth, nameSpace: nameSpace option) = displayValues(valu, depth, makeTypeEnv nameSpace) and propsForValue (Value {locations, typeOf, ...}) = PTtype typeOf :: mapLocationProps locations fun isConstructor (Value{class = Exception, ...}) = true | isConstructor (Value{class = Constructor _, ...}) = true | isConstructor _ = false fun isException (Value{class = Exception, ...}) = true | isException _ = false in val () = applyList (fn (name, v, t) => #enterVal valuesEnv (name, mkGvar (name, t, mkConst v, declInBasis))) [ ("name", toMachineWord (valName: values -> string), valueVal ->> String), ("print", toMachineWord (printValues: values * FixedInt.int -> pretty), mkProductType[valueVal, fixedIntType] ->> PrettyType), ("printWithType", toMachineWord (exportedDisplayValues: values * FixedInt.int * nameSpace option -> pretty), mkProductType[valueVal, fixedIntType, Option nameSpaceType] ->> PrettyType), ("printType", toMachineWord(exportedDisplayTypeExp: types * FixedInt.int * nameSpace option -> pretty), mkProductType[Types, fixedIntType, Option nameSpaceType] ->> PrettyType), ("typeof", toMachineWord (valTypeOf: values -> types), valueVal ->> Types), ("code", toMachineWord (codeForValue: values -> codetree), valueVal ->> CodeTree), ("properties", toMachineWord (propsForValue: values ->ptProperties list), valueVal ->> List PtProperties), ("isConstructor", toMachineWord(isConstructor: values -> bool), valueVal ->> Bool), ("isException", toMachineWord(isException: values -> bool), valueVal ->> Bool) ] end local (* TypeConstrs substructure. *) fun exportedDisplayTypeConstr(tyCons, depth, nameSpace: nameSpace option) = TYPETREE.displayTypeConstrs(tyCons, depth, makeTypeEnv nameSpace) and propsForTypeConstr (TypeConstrSet(TypeConstrs {locations, ...}, _)) = mapLocationProps locations and nameForType (TypeConstrSet(TypeConstrs{name, ...}, _)) = name in val () = applyList (fn (name, v, t) => #enterVal typesEnv (name, mkGvar (name, t, mkConst v, declInBasis))) [ ("name", toMachineWord(nameForType: typeConstrSet -> string), typeVal ->> String), ("print", toMachineWord (exportedDisplayTypeConstr: typeConstrSet * FixedInt.int * nameSpace option -> pretty), mkProductType[typeVal, fixedIntType, Option nameSpaceType] ->> PrettyType), ("properties", toMachineWord (propsForTypeConstr: typeConstrSet ->ptProperties list), typeVal ->> List PtProperties) ] end local (* Structures substructure *) fun exportedDisplayStructs(str, depth, nameSpace: nameSpace option) = displayStructures(str, depth, makeTypeEnv nameSpace) and codeForStruct (Struct{access = Global code, ...}) = code | codeForStruct _ = raise Fail "Not a global structure" and propsForStruct (Struct {locations, ...}) = mapLocationProps locations and nameForStruct (Struct{name, ...}) = name fun nameSpaceForStruct(baseStruct as Struct{signat=Signatures { tab, ...}, ...}): nameSpace = let open UNIVERSALTABLE fun lookupVal s = case univLookup (tab, valueVar, s) of NONE => NONE | SOME v => SOME(makeSelectedValue(v, baseStruct)) and lookupType s = case univLookup (tab, typeConstrVar, s) of NONE => NONE | SOME t => SOME(makeSelectedType(t, baseStruct)) and lookupStruct s = case univLookup (tab, structVar, s) of NONE => NONE | SOME s => SOME(makeSelectedStructure(s, baseStruct)) local fun extractItems t tab = UNIVERSALTABLE.fold (fn (s, u, l) => if Universal.tagIs t u then (s, Universal.tagProject t u) :: l else l ) [] tab in fun allValues() = map(fn (s, v) => (s, makeSelectedValue(v, baseStruct))) (extractItems valueVar tab) and allTypes() = map(fn (s, t) => (s, makeSelectedType(t, baseStruct))) (extractItems typeConstrVar tab) and allStructs() = map(fn (s, v) => (s, makeSelectedStructure(v, baseStruct))) (extractItems structVar tab) end fun enterFunction _ = raise Fail "updating a structure is not possible." (* Raise an exception for any attempt to enter a new value. Return empty for the classes that can't exist in a structure. *) in { lookupVal = lookupVal, lookupType = lookupType, lookupStruct = lookupStruct, lookupFix = fn _ => NONE, lookupSig = fn _ => NONE, lookupFunct = fn _ => NONE, enterVal = enterFunction, enterType = enterFunction, enterFix = enterFunction, enterStruct = enterFunction, enterSig = enterFunction, enterFunct = enterFunction, allVal = allValues, allType = allTypes, allStruct = allStructs, allFix = fn () => [], allSig = fn () => [], allFunct = fn () => [] } end in val () = applyList (fn (name, v, t) => #enterVal structsEnv (name, mkGvar (name, t, mkConst v, declInBasis))) [ ("name", toMachineWord(nameForStruct: structVals -> string), structureVal ->> String), ("print", toMachineWord (exportedDisplayStructs: structVals * FixedInt.int * nameSpace option -> pretty), mkProductType[structureVal, fixedIntType, Option nameSpaceType] ->> PrettyType), ("code", toMachineWord (codeForStruct: structVals -> codetree), structureVal ->> CodeTree), ("properties", toMachineWord (propsForStruct: structVals ->ptProperties list), structureVal ->> List PtProperties), ("contents", toMachineWord(nameSpaceForStruct: structVals -> nameSpace), structureVal ->> nameSpaceType) ] end local (* Signatures substructure *) fun exportedDisplaySigs(sign, depth, nameSpace: nameSpace option) = displaySignatures(sign, depth, makeTypeEnv nameSpace) and propsForSig (Signatures {locations, ...}) = mapLocationProps locations and nameForSig (Signatures{name, ...}) = name in val () = applyList (fn (name, v, t) => #enterVal sigsEnv (name, mkGvar (name, t, mkConst v, declInBasis))) [ ("name", toMachineWord(nameForSig: signatures -> string), signatureVal ->> String), ("print", toMachineWord (exportedDisplaySigs: signatures * FixedInt.int * nameSpace option -> pretty), mkProductType[signatureVal, fixedIntType, Option nameSpaceType] ->> PrettyType), ("properties", toMachineWord (propsForSig: signatures ->ptProperties list), signatureVal ->> List PtProperties) ] end local (* Functors substructure *) fun exportedDisplayFunctors(funct, depth, nameSpace: nameSpace option) = displayFunctors(funct, depth, makeTypeEnv nameSpace) and codeForFunct (Functor{access = Global code, ...}) = code | codeForFunct _ = raise Fail "Not a global functor" and propsForFunctor (Functor {locations, ...}) = mapLocationProps locations and nameForFunctor (Functor{name, ...}) = name in val () = applyList (fn (name, v, t) => #enterVal functsEnv (name, mkGvar (name, t, mkConst v, declInBasis))) [ ("name", toMachineWord(nameForFunctor: functors -> string), functorVal ->> String), ("print", toMachineWord (exportedDisplayFunctors: functors * FixedInt.int * nameSpace option -> pretty), mkProductType[functorVal, fixedIntType, Option nameSpaceType] ->> PrettyType), ("code", toMachineWord (codeForFunct: functors -> codetree), functorVal ->> CodeTree), ("properties", toMachineWord (propsForFunctor: functors ->ptProperties list), functorVal ->> List PtProperties) ] end local (* Infixes substructure *) fun nameForFix(FixStatus(s, _)) = s in val () = applyList (fn (name, v, t) => #enterVal fixesEnv (name, mkGvar (name, t, mkConst v, declInBasis))) [ ("name", toMachineWord(nameForFix: fixStatus -> string), fixityVal ->> String), ("print", toMachineWord (displayFixStatus: fixStatus -> pretty), fixityVal ->> PrettyType) ] end in end in () end (* initGlobalEnv *); end; diff --git a/mlsource/MLCompiler/TYPEIDCODE.sml b/mlsource/MLCompiler/TYPEIDCODE.sml index 4193fdad..c8f51ed2 100644 --- a/mlsource/MLCompiler/TYPEIDCODE.sml +++ b/mlsource/MLCompiler/TYPEIDCODE.sml @@ -1,1374 +1,1375 @@ (* - Copyright (c) 2009, 2013, 2015-16 David C. J. Matthews + Copyright (c) 2009, 2013, 2015-16, 2020 David C. J. Matthews This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) functor TYPEIDCODE ( structure LEX : LEXSIG; structure CODETREE : CODETREESIG structure TYPETREE : TYPETREESIG structure STRUCTVALS : STRUCTVALSIG structure DEBUG: DEBUGSIG structure PRETTY : PRETTYSIG structure ADDRESS : AddressSig sharing LEX.Sharing = STRUCTVALS.Sharing = PRETTY.Sharing = CODETREE.Sharing = TYPETREE.Sharing = ADDRESS ) : TYPEIDCODESIG = struct open CODETREE PRETTY ADDRESS STRUCTVALS TYPETREE (* This module deals with handling the run-time values that carry type information. At the moment that's just the equality and print operations but that will be extended. There are different versions according to whether this is a monomorphic constructor, a polymorphic constructor or a type. Monomorphic and polymorphic constructor values are passed around in the module system as run-time values for types and datatypes whereas type values are passed in the core language as an extra argument to polymorphic functions. Both monomorphic and polymorphic constructors contain a reference for the "printer" entry so that a pretty printer can be installed. The functions in polymorphic datatypes have to be applied to type values for the base types to construct a type value. Monomorphic datatypes just need some transformation. The effective types in each case are PolyType : (T('a) -> <'a t, 'a t> -> bool) * (T('a) -> 'a t * int -> pretty) ref MonoType : ( -> bool) * (t * int -> pretty) ref Type: ( -> bool) * (t * int -> pretty) where < > denotes multiple (poly-style) arguments rather than tuples. *) (* If this is true we are just using additional arguments for equality type variables. If false we are using them for all type variables and every polymorphic function is wrapped in a function that passes the type information. *) val justForEqualityTypes = true val arg1 = mkLoadArgument 0 (* Used frequently. *) val arg2 = mkLoadArgument 1 val InternalError = Misc.InternalError val orb = Word8.orb infix 7 orb; val mutableFlags = F_words orb F_mutable (* codeAccess is copied from ValueOps. *) fun codeAccess (Global code, _) = code | codeAccess (Local{addr=ref locAddr, level=ref locLevel}, level) = mkLoad (locAddr, level, locLevel) | codeAccess (Selected{addr, base}, level) = mkInd (addr, codeAccess (base, level)) | codeAccess _ = raise InternalError "No access" (* Load an identifier. *) fun codeId(TypeId{access, ...}, level) = codeAccess(access, level) (* Pretty printer code. These produce code to apply the pretty printer functions. *) fun codePrettyString(s: string) = mkDatatype[mkConst(toMachineWord tagPrettyString), mkConst(toMachineWord s)] and codePrettyBreak(n, m) = mkDatatype[mkConst(toMachineWord tagPrettyBreak), mkConst(toMachineWord n), mkConst(toMachineWord m)] and codePrettyBlock(n: int, t: bool, c: context list, args: codetree) = mkDatatype[mkConst(toMachineWord tagPrettyBlock), mkConst(toMachineWord n), mkConst(toMachineWord t), mkConst(toMachineWord c), args] (* Turn a list of codetrees into a run-time list. *) and codeList(c: codetree list, tail: codetree): codetree = List.foldr (fn (hd, tl) => mkTuple[hd, tl]) tail c (* Generate code to check that the depth is not less than the allowedDepth and if it is to print "..." rather than the given code. *) and checkDepth(depthCode: codetree, allowedDepth: int, codeOk, codeFail) = mkIf(mkBinary(BuiltIns.WordComparison{test=BuiltIns.TestLess, isSigned=true}, depthCode, mkConst(toMachineWord allowedDepth)), codeFail, codeOk) (* Subtract one from the current depth to produce the depth for sub-elements. *) and decDepth depthCode = mkBinary(BuiltIns.FixedPrecisionArith BuiltIns.ArithSub, depthCode, mkConst(toMachineWord 1)) val codePrintDefault = mkProc(codePrettyString "?", 1, "print-default", [], 0) structure TypeVarMap = struct (* Entries are either type var maps or "stoppers". *) datatype typeVarMapEntry = TypeVarFormEntry of (typeVarForm * (level->codetree)) list | TypeConstrListEntry of typeConstrs list type typeVarMap = { entryType: typeVarMapEntry, (* Either the type var map or a "stopper". *) cache: (* Cache of new type values. *) {typeOf: types, address: int, decCode: codeBinding} list ref, mkAddr: int->int, (* Make new addresses at this level. *) level: level (* Function nesting level. *) } list (* Default map. *) fun defaultTypeVarMap (mkAddr, level) = [{entryType=TypeConstrListEntry[], cache=ref [], mkAddr=mkAddr, level=level}] fun markTypeConstructors(typConstrs, mkAddr, level, tvs) = {entryType = TypeConstrListEntry typConstrs, cache = ref [], mkAddr=mkAddr, level=level} :: tvs fun getCachedTypeValues(({cache=ref cached, ...}) ::_): codeBinding list = (* Extract the values from the list. The later values may refer to earlier so the list must be reversed. *) List.rev (List.map (fn{decCode, ...} => decCode) cached) | getCachedTypeValues _ = raise Misc.InternalError "getCachedTypeValues" (* Extend a type variable environment with a new map of type variables to load functions. *) fun extendTypeVarMap (tvMap: (typeVarForm * (level->codetree)) list, mkAddr, level, typeVarMap) = {entryType = TypeVarFormEntry tvMap, cache = ref [], mkAddr=mkAddr, level=level} :: typeVarMap (* If we find the type var in the map return it as a type. This is used to eliminate apparently generalisable type vars from the list. *) fun mapTypeVars [] _ = NONE | mapTypeVars ({entryType=TypeVarFormEntry typeVarMap, ...} :: rest) tyVar = ( case List.find(fn(t, _) => sameTv(t, tyVar)) typeVarMap of SOME (tv, _) => SOME(TypeVar tv) | NONE => mapTypeVars rest tyVar ) | mapTypeVars (_ :: rest) tyVar = mapTypeVars rest tyVar (* Check to see if a type constructor is in the "stopper" set and return the level if it is. *) fun checkTypeConstructor(_, []) = ~1 (* Not there. *) | checkTypeConstructor(tyCons, {entryType=TypeVarFormEntry _, ...} :: rest) = checkTypeConstructor(tyCons, rest: typeVarMap) | checkTypeConstructor(tyCons, {entryType=TypeConstrListEntry tConstrs, ...} :: rest) = if List.exists(fn t => sameTypeId(tcIdentifier t, tcIdentifier tyCons)) tConstrs then List.length rest + 1 else checkTypeConstructor(tyCons, rest) local open TypeValue (* The printer and equality functions must be valid functions even when they will never be called. We may have to construct dummy type values by applying a polymorphic type constructor to them and if they don't have the right form the optimiser will complain. If we're only using type values for equality type variables the default print function will be used in polymorphic functions so must print "?". *) val errorFunction2 = mkProc(CodeZero, 2, "errorCode2", [], 0) val codeFn = mkProc(codePrettyString "fn", 1, "print-function", [], 0) local fun typeValForMonotype typConstr = let val codedId = codeId(tcIdentifier typConstr, baseLevel) val printerRefAddress = extractPrinter codedId val printFn = (* Create a function to load the printer ref and apply to the args. *) mkProc( mkEval( mkLoadOperation(LoadStoreMLWord{isImmutable=false}, printerRefAddress, CodeZero), [arg1]), 1, "print-" ^ tcName typConstr, [], 0) in createTypeValue{ eqCode=extractEquality codedId, printCode=printFn, boxedCode=extractBoxed codedId, sizeCode=extractSize codedId} end in (* A few common types. These are effectively always cached. *) val fixedIntCode = typeValForMonotype fixedIntConstr and intInfCode = typeValForMonotype intInfConstr and boolCode = typeValForMonotype boolConstr and stringCode = typeValForMonotype stringConstr and charCode = typeValForMonotype charConstr end (* Code generate this now so we only get one entry. *) val codeTuple = mkTuple[ createTypeValue{ (* Unused type variable. *) eqCode=errorFunction2, printCode=codePrintDefault, boxedCode=boxedEither, sizeCode=singleWord}, createTypeValue{ (* Function. *) eqCode=errorFunction2, printCode=codeFn, boxedCode=boxedAlways, sizeCode=singleWord}, fixedIntCode, intInfCode, boolCode, stringCode, charCode ] val code = genCode(codeTuple, [], 0)() in (* Default code used for a type variable that is not referenced but needs to be provided to satisfy the type. *) val defaultTypeCode = mkInd(0, code) val functionCode = mkInd(1, code) val cachedCode = [(fixedIntConstr, mkInd(2, code)), (intInfConstr, mkInd(3, code)), (boolConstr, mkInd(4, code)), (stringConstr, mkInd(5, code)), (charConstr, mkInd(6, code))] end fun findCachedTypeCode(typeVarMap: typeVarMap, typ): ((level->codetree) * int) option = let (* Test if we have the same type as the cached type. *) fun sameType (t1, t2) = case (eventual t1, eventual t2) of (TypeVar tv1, TypeVar tv2) => ( case (tvValue tv1, tvValue tv2) of (EmptyType, EmptyType) => sameTv(tv1, tv2) | _ => false ) | (FunctionType{arg=arg1, result=result1}, FunctionType{arg=arg2, result=result2}) => sameType(arg1, arg2) andalso sameType(result1, result2) | (LabelledType{recList=list1, ...}, LabelledType{recList=list2, ...}) => ListPair.allEq( fn({name=n1, typeof=t1}, {name=n2, typeof=t2}) => n1 = n2 andalso sameType(t1, t2)) (list1, list2) | (TypeConstruction{constr=c1, args=a1, ...}, TypeConstruction{constr=c2, args=a2, ...}) => sameTypeConstr(c1, c2) andalso ListPair.allEq sameType (a1, a2) | _ => false and sameTypeConstr(tc1, tc2) = sameTypeId(tcIdentifier tc1, tcIdentifier tc2) fun findCodeFromCache([], _) = NONE | findCodeFromCache(({cache=ref cache, level, ...} :: rest): typeVarMap, ty) = ( case List.find(fn {typeOf, ...} => sameType(typeOf, ty)) cache of NONE => findCodeFromCache(rest, ty) | SOME{address, ...} => SOME(fn l => mkLoad(address, l, level), List.length rest +1) ) in case typ of TypeVar tyVar => ( case tvValue tyVar of EmptyType => let (* If it's a type var it is either in the type var list or we return the default. It isn't in the cache. *) fun findCodeFromTypeVar([], _) = ((fn _ => defaultTypeCode), 0) (* Return default code for a missing type variable. This can occur if we have unreferenced type variables that need to be supplied but are treated as "don't care". *) | findCodeFromTypeVar({entryType=TypeVarFormEntry typeVarMap, ...} :: rest, tyVar) = ( case List.find(fn(t, _) => sameTv(t, tyVar)) typeVarMap of SOME(_, codeFn) => (codeFn, List.length rest+1) | NONE => findCodeFromTypeVar(rest, tyVar) ) | findCodeFromTypeVar(_ :: rest, tyVar) = findCodeFromTypeVar(rest, tyVar) in SOME(findCodeFromTypeVar(typeVarMap, tyVar)) end | OverloadSet _ => let val constr = typeConstrFromOverload(typ, false) in findCachedTypeCode(typeVarMap, mkTypeConstruction(tcName constr, constr, [], [])) end | ty => findCachedTypeCode(typeVarMap, ty) ) | TypeConstruction { constr, args, ...} => let fun sameTypeConstr(tc1, tc2) = sameTypeId(tcIdentifier tc1, tcIdentifier tc2) in if tcIsAbbreviation constr (* Type abbreviation *) then findCachedTypeCode(typeVarMap, makeEquivalent (constr, args)) else if null args then (* Check the permanently cached monotypes. *) case List.find(fn (t, _) => sameTypeConstr(t, constr)) cachedCode of SOME (_, c) => SOME ((fn _ => c), ~1) | NONE => findCodeFromCache(typeVarMap, typ) else findCodeFromCache(typeVarMap, typ) end | FunctionType _ => SOME(fn _ => functionCode, ~1) (* Every function has the same code. *) | _ => findCodeFromCache(typeVarMap, typ) end end open TypeVarMap (* Find the earliest entry in the cache table where we can put this entry. *) fun getMaxDepth (typeVarMap: typeVarMap) (ty: types, maxSoFar:int) : int = case findCachedTypeCode(typeVarMap, ty) of SOME (_, cacheDepth) => Int.max(cacheDepth, maxSoFar) | NONE => let in case ty of TypeVar tyVar => ( case tvValue tyVar of OverloadSet _ => maxSoFar (* Overloads are all global. *) | EmptyType => maxSoFar | tyVal => getMaxDepth typeVarMap (tyVal, maxSoFar) ) | TypeConstruction{constr, args, ...} => if tcIsAbbreviation constr (* May be an alias *) then getMaxDepth typeVarMap (makeEquivalent (constr, args), maxSoFar) else List.foldl (getMaxDepth typeVarMap) (Int.max(maxSoFar, checkTypeConstructor(constr, typeVarMap))) args | LabelledType {recList, ...} => List.foldl (fn ({typeof, ...}, m) => getMaxDepth typeVarMap (typeof, m)) maxSoFar recList | _ => maxSoFar end (* Get the boxedness status for a type i.e. whether values of the type are always addresses, always tagged integers or could be either. *) fun boxednessForType(ty, level: level, getTypeValueForID, typeVarMap): codetree = case findCachedTypeCode(typeVarMap, ty) of SOME (code, _) => TypeValue.extractBoxed(code level) | NONE => let fun boxednessForConstruction(constr, args): codetree = (* Get the boxedness for a datatype construction. *) let (* Get the boxedness functions for the argument types. This applies only to polytypes. *) fun getArg ty : codetree = let val boxedFun = boxednessForType(ty, level, getTypeValueForID, typeVarMap) open TypeValue in (* We need a type value here although only the boxedFun will be used. *) createTypeValue{eqCode=CodeZero, printCode=CodeZero, boxedCode=boxedFun, sizeCode=singleWord} end val codeForId = TypeValue.extractBoxed(getTypeValueForID(tcIdentifier constr, args, level)) in (* Apply the function we obtained to any type arguments. *) if null args then codeForId else mkEval(codeForId, map getArg args) end in case ty of TypeVar tyVar => ( case tvValue tyVar of OverloadSet _ => boxednessForConstruction(typeConstrFromOverload(ty, false), []) | EmptyType => raise InternalError "boxedness: should already have been handled" | tyVal => boxednessForType(tyVal, level, getTypeValueForID, typeVarMap) ) | TypeConstruction{constr, args, ...} => if tcIsAbbreviation constr (* May be an alias *) then boxednessForType (makeEquivalent (constr, args), level, getTypeValueForID, typeVarMap) else boxednessForConstruction(constr, args) | LabelledType {recList=[{typeof=singleton, ...}], ...} => (* Unary tuples are optimised - no indirection. *) boxednessForType(singleton, level, getTypeValueForID, typeVarMap) | LabelledType _ => TypeValue.boxedAlways (* Tuple are currently always boxed. *) (* Functions are handled in the cache case. *) | _ => raise InternalError "boxednessForType: Unknown type" end (* Get the size for values of the type. A value N other than 1 means that every value of the type is a pointer to a tuple of exactly N words. Zero is never used. *) fun sizeForType(ty, level, getTypeValueForID, typeVarMap): codetree = case findCachedTypeCode(typeVarMap, ty) of SOME (code, _) => TypeValue.extractSize(code level) | NONE => let fun sizeForConstruction(constr, args): codetree = (* Get the size for a datatype construction. *) let (* Get the size functions for the argument types. This applies only to polytypes. *) fun getArg ty : codetree = let val sizeFun = sizeForType(ty, level, getTypeValueForID, typeVarMap) open TypeValue in (* We need a type value here although only the sizeFun will be used. *) createTypeValue{eqCode=CodeZero, printCode=CodeZero, boxedCode=CodeZero, sizeCode=sizeFun} end val codeForId = TypeValue.extractSize(getTypeValueForID(tcIdentifier constr, args, level)) in (* Apply the function we obtained to any type arguments. *) if null args then codeForId else mkEval(codeForId, map getArg args) end in case ty of TypeVar tyVar => ( case tvValue tyVar of OverloadSet _ => sizeForConstruction(typeConstrFromOverload(ty, false), []) | EmptyType => raise InternalError "size: should already have been handled" | tyVal => sizeForType(tyVal, level, getTypeValueForID, typeVarMap) ) | TypeConstruction{constr, args, ...} => if tcIsAbbreviation constr (* May be an alias *) then sizeForType (makeEquivalent (constr, args), level, getTypeValueForID, typeVarMap) else sizeForConstruction(constr, args) | LabelledType {recList=[{typeof=singleton, ...}], ...} => (* Unary tuples are optimised - no indirection. *) sizeForType(singleton, level, getTypeValueForID, typeVarMap) | LabelledType{recList, ...} => let val length = List.length recList in (* Set the length to the number of words that can be unpacked. If there are more than 4 items it's probably not worth packing them into other tuples so set this to one. *) if length <= 4 (*!maxPacking*) then mkConst(toMachineWord length) else TypeValue.singleWord end (* Functions are handled in the cache case. *) | _ => raise InternalError "sizeForType: Unknown type" end fun printerForType(ty, baseLevel, argTypes: typeVarMap) = let fun printCode(typ, level: level) = ( case typ of typ as TypeVar tyVar => ( case tvValue tyVar of EmptyType => ( case findCachedTypeCode(argTypes, typ) of SOME (code, _) => TypeValue.extractPrinter(code level) | NONE => raise InternalError "printerForType: should already have been handled" ) | OverloadSet _ => let val constr = typeConstrFromOverload(typ, false) in printCode(mkTypeConstruction(tcName constr, constr, [], []), level) end | _ => (* Just a bound type variable. *) printCode(tvValue tyVar, level) ) | TypeConstruction { constr=typConstr, args, name, ...} => if tcIsAbbreviation typConstr (* Handle type abbreviations directly *) then printCode(makeEquivalent (typConstr, args), level) else let val nLevel = newLevel level (* Get the type Id and put in code to extract the printer ref. *) val codedId = codeId(tcIdentifier typConstr, nLevel) open TypeValue val printerRefAddress = extractPrinter codedId (* We need a type value here. The printer field will be used to print the type argument and the boxedness and size fields may be needed to extract the argument from the constructed value. *) fun makePrinterId t = let fun codeForId(typeId, _, l) = codeId(typeId, l) in createTypeValue {eqCode=CodeZero, printCode=printCode(t, nLevel), boxedCode=boxednessForType(t, nLevel, codeForId, argTypes), sizeCode=sizeForType(t, nLevel, codeForId, argTypes)} end val argList = map makePrinterId args in case args of [] => (* Create a function that, when called, will extract the function from the reference and apply it the pair of the value and the depth. *) mkProc( mkEval( mkLoadOperation(LoadStoreMLWord{isImmutable=false}, printerRefAddress, CodeZero), [arg1]), 1, "print-"^name, getClosure nLevel, 0) | _ => (* Construct a function, that when called, will extract the function from the reference and apply it first to the base printer functions and then to the pair of the value and depth. *) mkProc( mkEval( mkEval( mkLoadOperation(LoadStoreMLWord{isImmutable=false}, printerRefAddress, CodeZero), argList), [arg1]), 1, "print-"^name, getClosure nLevel, 0) end | LabelledType { recList=[], ...} => (* Empty tuple: This is the unit value. *) mkProc(codePrettyString "()", 1, "print-labelled", [], 0) | LabelledType {recList=[{name, typeof}], ...} => let (* Optimised unary record *) val localLevel = newLevel level val entryCode = mkEval(printCode(typeof, localLevel), [arg1]) val printItem = codeList([codePrettyString(name^" ="), codePrettyBreak(1, 0), entryCode, codePrettyString "}"], CodeZero) in mkProc( codePrettyBlock(1, false, [], mkTuple[codePrettyString "{", printItem]), 1, "print-labelled", getClosure localLevel, 0) end | LabelledType (r as { recList, ...}) => let (* See if this has fields numbered 1=, 2= etc. N.B. If it has only one field we need to print 1= since we don't have singleton tuples. *) fun isRec([], _) = true | isRec({name, ...} :: l, n) = name = Int.toString n andalso isRec(l, n+1) val isTuple = recordIsFrozen r andalso isRec(recList, 1) andalso List.length recList >= 2 val localLevel = newLevel level val valToPrint = mkInd(0, arg1) and depthCode = mkInd(1, arg1) val fields = List.tabulate(List.length recList, fn n => n) val items = ListPair.zipEq(recList, fields) (* The ordering on fields is designed to allow mixing of tuples and records (e.g. #1). It puts shorter names before longer so that #11 comes after #2 and before #100. For named records it does not make for easy reading so we sort those alphabetically when printing. *) val printItems = if isTuple then items else Misc.quickSort(fn ({name = a, ...}, _) => fn ({name = b, ...}, _) => a <= b) items fun asRecord([], _) = raise Empty (* Shouldn't happen. *) | asRecord([({name, typeof, ...}, offset)], _) = let val entryCode = (* Last field: no separator. *) mkEval(printCode(typeof, localLevel), [mkTuple[mkInd(offset, valToPrint), decDepth depthCode]]) val (start, terminator) = if isTuple then ([], ")") else ([codePrettyString(name^" ="), codePrettyBreak(1, 0)], "}") in codeList(start @ [entryCode, codePrettyString terminator], CodeZero) end | asRecord(({name, typeof, ...}, offset) :: fields, depth) = let val (start, terminator) = if isTuple then ([], ")") else ([codePrettyString(name^" ="), codePrettyBreak(1, 0)], "}") in checkDepth(depthCode, depth, codeList( start @ [ mkEval( printCode(typeof, localLevel), [mkTuple[mkInd(offset, valToPrint), decDepth depthCode]]), codePrettyString ",", codePrettyBreak (1, 0) ], asRecord(fields, depth+1)), codeList([codePrettyString ("..." ^ terminator)], CodeZero) ) end in mkProc( codePrettyBlock(1, false, [], mkTuple[codePrettyString (if isTuple then "(" else "{"), asRecord(printItems, 0)]), 1, "print-labelled", getClosure localLevel, 0) end | FunctionType _ => mkProc(codePrettyString "fn", 1, "print-function", [], 0) | _ => mkProc(codePrettyString "", 1, "print-empty", [], 0) ) in printCode(ty, baseLevel) end and makeEq(ty, level: level, getTypeValueForID, typeVarMap): codetree = let fun equalityForConstruction(constr, args): codetree = (* Generate an equality function for a datatype construction. *) let (* Get argument types parameters for polytypes. There's a special case here for type vars, essentially the type arguments to the datatype, to avoid taking apart the type value record and then building it again. *) fun getArg ty = if (case ty of TypeVar tyVar => (case tvValue tyVar of EmptyType => true | _ => false) | _ => false) then ( case findCachedTypeCode(typeVarMap, ty) of SOME (code, _) => code level | NONE => raise InternalError "getArg" ) else let val eqFun = makeEq(ty, level, getTypeValueForID, typeVarMap) open TypeValue in (* We need a type value here. The equality function will be used to compare the argument type and the boxedness and size parameters may be needed for the constructors. *) createTypeValue{eqCode=eqFun, printCode=CodeZero, boxedCode=boxednessForType(ty, level, getTypeValueForID, typeVarMap), sizeCode=sizeForType(ty, level, getTypeValueForID, typeVarMap)} end val resFun = let val iden = tcIdentifier constr in (* Special case: If this is ref, Array.array or Array2.array we must use pointer equality and not attempt to create equality functions for the argument. It may not be an equality type. *) if isPointerEqType iden - then equalWordFn + then equalPointerOrWordFn else let open TypeValue val codeForId = extractEquality(getTypeValueForID(tcIdentifier constr, args, level)) in (* Apply the function we obtained to any type arguments. *) if null args then codeForId else mkEval(codeForId, map getArg args) end end in resFun end in case ty of TypeVar tyVar => ( case tvValue tyVar of OverloadSet _ => (* This seems to occur if there are what amount to indirect references to literals. *) equalityForConstruction(typeConstrFromOverload(ty, false), []) | EmptyType => ( case findCachedTypeCode(typeVarMap, ty) of SOME (code, _) => TypeValue.extractEquality(code level) | NONE => raise InternalError "makeEq: should already have been handled" ) | tyVal => makeEq(tyVal, level, getTypeValueForID, typeVarMap) ) | TypeConstruction{constr, args, ...} => if tcIsAbbreviation constr (* May be an alias *) then makeEq (makeEquivalent (constr, args), level, getTypeValueForID, typeVarMap) else equalityForConstruction(constr, args) | LabelledType {recList=[{typeof=singleton, ...}], ...} => (* Unary tuples are optimised - no indirection. *) makeEq(singleton, level, getTypeValueForID, typeVarMap) | LabelledType {recList, ...} => (* Combine the entries. fun eq(a,b) = #1 a = #1 b andalso #2 a = #2 b ... *) let (* Have to turn this into a new function. *) val nLevel = newLevel level fun combineEntries ([], _) = CodeTrue | combineEntries ({typeof, ...} :: t, n) = let val compareElements = makeEq(typeof, nLevel, getTypeValueForID, typeVarMap) in mkCand( mkEval(compareElements, [mkInd(n, arg1), mkInd(n, arg2)]), combineEntries (t, n+1)) end val tupleCode = combineEntries(recList, 0) in mkProc(tupleCode, 2, "eq{...}(2)", getClosure nLevel, 0) end | _ => raise InternalError "Equality for function" end (* Create equality functions for a set of possibly mutually recursive datatypes. *) fun equalityForDatatypes(typeDataList, eqAddresses, baseEqLevel, typeVarMap): (int * codetree) list = let val typesAndAddresses = ListPair.zipEq(typeDataList, eqAddresses) fun equalityForDatatype(({typeConstr=TypeConstrSet(tyConstr, vConstrs), eqStatus, (*boxedCode, sizeCode,*) ...}, addr), otherFns) = if eqStatus then let val nTypeVars = tcArity tyConstr val argTypes = List.tabulate(tcArity tyConstr, fn _ => makeTv{value=EmptyType, level=generalisable, nonunifiable=false, equality=false, printable=false}) val baseEqLevelP1 = newLevel baseEqLevel (* Argument type variables. *) val (localArgList, argTypeMap) = case argTypes of [] => ([], typeVarMap) | _ => let (* Add the polymorphic variables after the ordinary ones. *) (* Create functions to load these if they are used in the map. They may be non-local!!! *) val args = List.tabulate(nTypeVars, fn addr => fn l => mkLoadParam(addr+2, l, baseEqLevelP1)) (* Put the outer args in the map *) val varToArgMap = ListPair.zipEq(argTypes, args) (* Load the local args to return. *) val localArgList = List.tabulate (nTypeVars, fn addr => mkLoadParam(addr+2, baseEqLevelP1, baseEqLevelP1)) val addrs = ref 0 (* Make local declarations for any type values. *) fun mkAddr n = !addrs before (addrs := !addrs + n) in (localArgList, extendTypeVarMap(varToArgMap, mkAddr, baseEqLevelP1, typeVarMap)) end (* If this is a reference to a datatype we're currently generating load that address otherwise fall back to the default. *) fun getEqFnForID(typeId, _, l) = (* if sameTypeId(typeId, tcIdentifier tyConstr) andalso null argTypes then (* Directly recursive. *) TypeValue.createTypeValue{eqCode=mkLoadRecursive(l-baseLevel-1), printCode=CodeZero, boxedCode=boxedCode, sizeCode=sizeCode} else *) case List.find(fn({typeConstr=tc, ...}, _) => sameTypeId(tcIdentifier(tsConstr tc), typeId)) typesAndAddresses of SOME({boxedCode, sizeCode, ...}, addr) => (* Mutually recursive. *) TypeValue.createTypeValue{eqCode=mkLoad(addr, l, baseEqLevel), printCode=CodeZero, boxedCode=boxedCode, sizeCode=sizeCode} | NONE => codeId(typeId, l) (* Filter out the ShortForm constructors. They arise in situations such as datatype t = A of int*int | B | C i.e. where we have only one non-nullary constructor and it is a tuple. In this case we can deal with all the nullary constructors simply by testing whether the two arguments are the same. We don't have to discriminate the individual cases. *) fun processConstrs [] = (* The last of the alternatives is false *) CodeZero | processConstrs (Value{class, access, typeOf, ...} :: rest) = let fun addPolymorphism c = if nTypeVars = 0 orelse justForEqualityTypes then c else mkEval(c, localArgList) val base = codeAccess(access, baseEqLevelP1) open ValueConstructor fun matches arg = mkEval(addPolymorphism(extractTest base), [arg]) in case class of Constructor{nullary=true, ...} => let (* Nullary constructors are represented either by short constants or by constant tuples depending on the rest of the datatype. If this is a short constant the pointer equality is sufficient. This appears to increase the code size but the test should be optimised away because it is applied to a constant. (The "injection function" of a nullary constructor is the constant that represents the value). We have to test the tags if it is not short because we can't guarantee that the constant tuple hasn't been duplicated. *) val isShort = mkIsShort(addPolymorphism(extractInjection base)) in mkIf(mkIf(isShort, CodeFalse, matches arg1), matches arg2, processConstrs rest) end | _ => (* We have to unwrap the value. *) let (* Get the constructor argument given the result type. We might actually be able to take the argument type off directly but there's some uncertainty about whether we use the same type variables for the constructors as for the datatype. (This only applies for polytypes). *) val resType = constructorResult(typeOf, List.map TypeVar argTypes) (* Code to extract the value. *) fun destruct argNo = mkEval(addPolymorphism(extractProjection(codeAccess(access, baseEqLevelP1))), [mkLoadParam(argNo, baseEqLevelP1, baseEqLevelP1)]) (* Test whether the values match. *) val eqValue = mkEval( makeEq(resType, baseEqLevelP1, getEqFnForID, argTypeMap), [destruct 0, destruct 1]) in (* We have equality if both values match this constructor and the values within the constructor match. *) mkIf(matches arg1, mkCand(matches arg2, eqValue), processConstrs rest) end end (* processConstrs assumes that if there are nullary constructors we have already tested for bitwise equality. We also do that if there is more than one constructor to try to speed up equality for deep structures. *) val eqCode = case vConstrs of [Value{class=Constructor{nullary=true, ...}, ...}] => CodeTrue | [_] => processConstrs vConstrs - | _ => mkCor(mkEqualWord(arg1, arg2), processConstrs vConstrs) + | _ => mkCor(mkEqualPointerOrWord(arg1, arg2), processConstrs vConstrs) in if null argTypes then (addr, mkProc(eqCode, 2, "eq-" ^ tcName tyConstr ^ "(2)", getClosure baseEqLevelP1, 0)) :: otherFns else (* Polymorphic. Add an extra inline functions. *) let val nArgs = List.length argTypes val nLevel = newLevel baseEqLevel val nnLevel = newLevel nLevel (* Call the second function with the values to be compared and the base types. *) val polyArgs = List.tabulate(nArgs, fn i => mkLoadParam(i, nnLevel, nLevel)) in (addr, mkInlproc( mkInlproc( mkEval(mkLoad(addr+1, nnLevel, baseEqLevel), [arg1, arg2] @ polyArgs), 2, "eq-" ^ tcName tyConstr ^ "(2)", getClosure nnLevel, 0), nArgs, "eq-" ^ tcName tyConstr ^ "(2)(P)", getClosure nLevel, 0)) :: (addr+1, mkProc(mkEnv(getCachedTypeValues argTypeMap, eqCode), 2+nTypeVars, "eq-" ^ tcName tyConstr ^ "()", getClosure baseEqLevelP1, 0)) :: otherFns end end else (* Not an equality type. This will not be called but it still needs to be a function to ensure it's valid inside mkMutualDecs. *) (addr, mkProc(CodeZero, 2, "no-eq", [], 0)) :: otherFns in List.foldl equalityForDatatype [] typesAndAddresses end (* Create a printer function for a datatype when the datatype is declared. We don't have to treat mutually recursive datatypes specially because this is called after the type IDs have been created. *) fun printerForDatatype(TypeConstrSet(typeCons as TypeConstrs{name, ...}, vConstrs), level, typeVarMap) = let val argCode = mkInd(0, arg1) and depthCode = mkInd(1, arg1) val nLevel = newLevel level val constrArity = tcArity typeCons val argTypes = List.tabulate(constrArity, fn _ => makeTv{value=EmptyType, level=generalisable, nonunifiable=false, equality=false, printable=false}) val (localArgList, innerLevel, newTypeVarMap) = case constrArity of 0 => ([], nLevel, typeVarMap) | _ => let val nnLevel = newLevel nLevel fun mkTcArgMap (argTypes, level, oldLevel) = let val nArgs = List.length argTypes val argAddrs = List.tabulate(nArgs, fn n => n) val args = List.map(fn addr => fn l => mkLoadParam(addr, l, oldLevel)) argAddrs in (ListPair.zipEq(argTypes, args), List.map (fn addr => mkLoadParam(addr, level, oldLevel)) argAddrs) end val (varToArgMap, localArgList) = mkTcArgMap(argTypes, nnLevel, nLevel) val addrs = ref 1 (* Make local declarations for any type values. *) fun mkAddr n = !addrs before (addrs := !addrs + n) in (localArgList, nnLevel, extendTypeVarMap(varToArgMap, mkAddr, nLevel, typeVarMap)) end (* If we have an expression as the argument we parenthesise it unless it is a simple string, a tuple, a record or a list. *) (* fun parenthesise p = let val test = case p of PrettyBlock(_, _, _, items) => ( case items of PrettyString first :: tl => not(null tl) andalso first <> "(" andalso first <> "{" andalso first <> "[" | _ => false ) | _ => false in if test then PrettyBlock(3, true, [], [ PrettyString "(", PrettyBreak(0, 0), p, PrettyBreak(0, 0), PrettyString ")" ]) else p end *) local - fun eqStr (arg, str) = mkEqualWord(arg, mkConst(toMachineWord str)) + fun eqStr (arg, str) = mkEqualPointerOrWord(arg, mkConst(toMachineWord str)) + (* eqStr assumes that all occurrences of the same single character string are shared. *) val isNotNull = mkNot o mkIsShort fun testTag(arg, tagV) = (* Test the tag in the first word of the datatype. *) mkTagTest(mkInd(0, arg), tagV, maxPrettyTag) fun listHd x = mkVarField(0, x) and listTl x = mkVarField(1, x) in val parenCode = mkProc( mkIf( testTag(mkLoadArgument 0, tagPrettyBlock), (* then *) mkEnv( [mkDec(0, mkVarField(4, mkLoadArgument 0))], (* items *) mkIf ( (* not(null items) andalso not(null(tl items)) andalso not (isPrettyString(hd items) andalso bracket) *) mkCand( isNotNull(mkLoadLocal 0), mkCand( isNotNull (listTl(mkLoadLocal 0)), mkNot ( mkCand(testTag(listHd(mkLoadLocal 0), tagPrettyString), mkEnv( [mkDec(1, mkVarField(1, listHd(mkLoadLocal 0)))], mkCor(eqStr(mkLoadLocal 1, "("), mkCor(eqStr(mkLoadLocal 1, "{"), eqStr(mkLoadLocal 1, "["))) ) ) ) ) ), (* then: Parenthesise the argument. *) codePrettyBlock( 3, true, [], mkDatatype [ codePrettyString "(", mkDatatype [ codePrettyBreak(0, 0), mkDatatype [ mkLoadArgument 0, mkDatatype [ codePrettyBreak(0, 0), mkDatatype [codePrettyString ")", CodeZero ] ] ] ] ] ), (* else *) mkLoadArgument 0 ) ), (* else *) mkLoadArgument 0 ), 1, "parenthesise", [], 2) end fun printerForConstructors (Value{name, typeOf, access, class = Constructor{nullary, ...}, locations, ...} :: rest) = let (* The "value" for a value constructor is a tuple containing the test code, the injection and the projection functions. *) val constructorCode = codeAccess(access, innerLevel) (* If this is a polytype the fields in the constructor tuple are functions that first have to be applied to the type arguments to yield the actual injection/test/projection functions. For monotypes the fields contain the injection/test/projection functions directly. *) fun addPolymorphism c = if constrArity = 0 orelse justForEqualityTypes then c else mkEval(c, localArgList) open ValueConstructor val locProps = (* Get the declaration location. *) List.foldl(fn (DeclaredAt loc, _) => [ContextLocation loc] | (_, l) => l) [] locations val nameCode = codePrettyBlock(0, false, locProps, codeList([codePrettyString name], CodeZero)) val printCode = if nullary then (* Just the name *) nameCode else let val typeOfArg = constructorResult(typeOf, List.map TypeVar argTypes) val getValue = mkEval(addPolymorphism(extractProjection constructorCode), [argCode]) in codePrettyBlock(1, false, [], codeList( [ (* Put it in a block with the declaration location. *) nameCode, codePrettyBreak (1, 0), (* Print the argument and parenthesise it if necessary. *) mkEval(parenCode, [ mkEval( printerForType(typeOfArg, innerLevel, newTypeVarMap), [mkTuple[getValue, decDepth depthCode]] )] ) ], CodeZero)) end in (* If this was the last or only constructor we don't need to test. *) checkDepth(depthCode, 1, if null rest then printCode else let val testValue = mkEval(addPolymorphism(extractTest constructorCode), [argCode]) in mkIf(testValue, printCode, printerForConstructors rest) end, codePrettyString "...") end | printerForConstructors _ = raise InternalError ("No constructors:"^name) val printerCode = printerForConstructors vConstrs in (* Wrap this in the functions for the base types. *) if constrArity = 0 then mkProc(printerCode, 1, "print-"^name, getClosure innerLevel, 0) else mkProc(mkEnv(getCachedTypeValues newTypeVarMap, mkProc(printerCode, 1, "print-"^name, getClosure innerLevel, 0)), constrArity, "print"^name^"()", getClosure nLevel, 0) end (* Opaque matching and functor application create new type IDs using an existing type as implementation. The equality function is inherited whether the type was specified as an eqtype or not. The print function is no longer inherited. Instead a new reference is installed with a default print function. This hides the implementation. *) (* If this is a type function we're going to generate a new ref anyway so we don't need to copy it. *) fun codeGenerativeId{source=TypeId{idKind=TypeFn([], resType), ...}, isEq, mkAddr, level, ...} = let (* Monotype abbreviation. *) (* Create a new type value cache. *) val typeVarMap = defaultTypeVarMap(mkAddr, level) open TypeValue val eqCode = if not isEq then CodeZero else (* We need a function that takes two arguments rather than a single pair. *) makeEq(resType, level, fn (typeId, _, l) => codeId(typeId, l), typeVarMap) val boxedCode = boxednessForType(resType, level, fn (typeId, _, l) => codeId(typeId, l), typeVarMap) val sizeCode = sizeForType(resType, level, fn (typeId, _, l) => codeId(typeId, l), typeVarMap) in mkEnv( TypeVarMap.getCachedTypeValues typeVarMap, createTypeValue { eqCode = eqCode, boxedCode = boxedCode, sizeCode = sizeCode, printCode = mkAllocateWordMemory( mkConst (toMachineWord 1), mkConst (toMachineWord mutableFlags), codePrintDefault) }) end | codeGenerativeId{source=TypeId{idKind=TypeFn(argTypes, resType), ...}, isEq, mkAddr, level, ...} = let (* Polytype abbreviation: All the entries in the tuple are functions that must be applied to the base type values when the type constructor is used. *) (* Create a new type value cache. *) val typeVarMap = defaultTypeVarMap(mkAddr, level) val nArgs = List.length argTypes fun createCode(makeCode, name) = let val nLevel = newLevel level val addrs = ref 0 fun mkAddr n = !addrs before (addrs := !addrs + n) local val args = List.tabulate(nArgs, fn addr => fn l => mkLoadParam(addr, l, nLevel)) in val typeEnv = ListPair.zipEq(argTypes, args) end val argTypeMap = extendTypeVarMap(typeEnv, mkAddr, nLevel, typeVarMap) val innerFnCode = makeCode(nLevel, argTypeMap) in mkProc(mkEnv(getCachedTypeValues argTypeMap, innerFnCode), nArgs, name, getClosure nLevel, !addrs) end open TypeValue (* Create a print function.*) val printCode = createCode(fn _ => codePrintDefault, "print-helper()") and eqCode = if not isEq then CodeZero else createCode(fn(nLevel, argTypeMap) => makeEq(resType, nLevel, fn (typeId, _, l) => codeId(typeId, l), argTypeMap), "equality()") and boxedCode = createCode(fn(nLevel, argTypeMap) => boxednessForType(resType, nLevel, fn (typeId, _, l) => codeId(typeId, l), argTypeMap), "boxedness()") and sizeCode = createCode(fn(nLevel, argTypeMap) => sizeForType(resType, nLevel, fn (typeId, _, l) => codeId(typeId, l), argTypeMap), "size()") in mkEnv( TypeVarMap.getCachedTypeValues typeVarMap, createTypeValue { eqCode = eqCode, boxedCode = boxedCode, printCode = mkAllocateWordMemory( mkConst (toMachineWord 1), mkConst (toMachineWord mutableFlags), printCode), sizeCode = sizeCode }) end | codeGenerativeId{source=sourceId, isDatatype, mkAddr, level, ...} = let (* Datatype. This is the same for monotype and polytypes except for the print fn. *) (* We hide the print function if the target is just a type name but if the target is a datatype it's probably better to have a print function. We inherit it from the source although that may expose the representation of other types. e.g. structure S:> sig type t datatype s = A of t end = ... *) open TypeValue val { dec, load } = multipleUses (codeId(sourceId, level), fn () => mkAddr 1, level) val loadLocal = load level val arity = case sourceId of TypeId{idKind=Bound{arity, ...},...} => arity | TypeId{idKind=Free{arity, ...},...} => arity | TypeId{idKind=TypeFn _,...} => raise InternalError "Already checked" val printFn = if isDatatype then mkLoadOperation(LoadStoreMLWord{isImmutable=false}, extractPrinter loadLocal, CodeZero) else if arity = 0 then codePrintDefault else mkProc(codePrintDefault, arity, "print-helper()", [], 0) val printCode = mkAllocateWordMemory( mkConst (toMachineWord 1), mkConst (toMachineWord mutableFlags), printFn) in mkEnv( dec, createTypeValue { eqCode = extractEquality loadLocal, printCode = printCode, boxedCode = extractBoxed loadLocal, sizeCode = extractSize loadLocal } ) end (* Create the equality and type functions for a set of mutually recursive datatypes. *) fun createDatatypeFunctions( typeDatalist: {typeConstr: typeConstrSet, eqStatus: bool, boxedCode: codetree, sizeCode: codetree } list, mkAddr, level, typeVarMap, makePrintFunction) = let (* Each entry has an equality function and a ref to a print function. The print functions for each type needs to indirect through the refs when printing other types so that if a pretty printer is later installed for one of the types the others will use the new pretty printer. That means that the code has to be produced in stages. *) (* Create the equality functions. Because mutual decs can only be functions we can't create the typeIDs themselves as mutual declarations. *) local (* If this is polymorphic make two addresses, one for the returned equality function and one for the inner function. *) fun makeEqAddr{typeConstr=TypeConstrSet(tyConstr, _), ...} = mkAddr(if tcArity tyConstr = 0 then 1 else 2) in val eqAddresses = List.map makeEqAddr typeDatalist (* Make addresses for the equalities. *) end val equalityFunctions = mkMutualDecs(equalityForDatatypes(typeDatalist, eqAddresses, level, typeVarMap)) (* Create the typeId values and set their addresses. The print function is initially set as zero. *) local fun makeTypeId({typeConstr, boxedCode, sizeCode, ...}, eqAddr) = let val var = vaLocal(idAccess(tcIdentifier(tsConstr typeConstr))) val newAddr = mkAddr 1 open TypeValue val idCode = createTypeValue { eqCode=mkLoadLocal eqAddr, printCode= mkAllocateWordMemory( mkConst (toMachineWord 1), mkConst (toMachineWord mutableFlags), CodeZero (* Temporary - replaced by setPrinter. *)), boxedCode = boxedCode, sizeCode = sizeCode } in #addr var := newAddr; #level var:= level; mkDec(newAddr, idCode) end in val typeIdCode = ListPair.map makeTypeId (typeDatalist, eqAddresses) end (* Create the print functions and set the printer code for each typeId. *) local fun setPrinter{typeConstr as TypeConstrSet(tCons as TypeConstrs{identifier, ...}, _), ...} = let val arity = tcArity tCons val printCode = if makePrintFunction then printerForDatatype(typeConstr, level, typeVarMap) else if arity = 0 then codePrintDefault else mkProc(codePrintDefault, arity, "print-printdefault", [], 0) in mkNullDec( mkStoreOperation(LoadStoreMLWord{isImmutable=false}, TypeValue.extractPrinter(codeId(identifier, level)), CodeZero, printCode)) end in val printerCode = List.map setPrinter typeDatalist end in equalityFunctions :: typeIdCode @ printerCode end (* Exported function. Returns a function from an ML pair of values to bool. N.B. This differs from the functions in the typeID which take a Poly pair. *) fun equalityForType(ty: types, level: level, typeVarMap: typeVarMap): codetree = let val nLevel = newLevel level (* The final result function must take a single argument. *) val resultCode = makeEq(ty, nLevel, fn (typeId, _, l) => codeId(typeId, l), typeVarMap) in (* We need to wrap this up in a new inline function. *) mkInlproc(mkEval(resultCode, [mkInd(0, arg1), mkInd(1, arg1)]), 1, "equality", getClosure nLevel, 0) end (* This code is used when the type checker has to construct a unique monotype because a type variable has escaped to the top level. The equality code always returns true and the printer prints "?". *) fun codeForUniqueId() = let open TypeValue val alwaysTrue = mkProc(CodeTrue, 2, "codeForUniqueId-equal", [], 0) val printCode = mkAllocateWordMemory( mkConst (toMachineWord 1), mkConst (toMachineWord mutableFlags), codePrintDefault) in createTypeValue{ eqCode = alwaysTrue, printCode = printCode, boxedCode = boxedEither, sizeCode = singleWord } end val noEquality = mkProc(CodeFalse, 2, "noEquality", [], 0) (* Since we don't have a way of writing a "printity" type variable there are cases when the printer will have to fall back to this. e.g. if we have a polymorphic printing function as a functor argument. *) val noPrinter = codePrintDefault (* If this is a polymorphic value apply it to the type instance. *) fun applyToInstance'([], level, _, code) = code level (* Monomorphic. *) | applyToInstance'(sourceTypes, level, polyVarMap, code) = let (* If we need either the equality or print function we generate a new entry and ignore anything in the cache. *) fun makePolyParameter {value=t, equality, printity} = if equality orelse printity then let open TypeValue fun getTypeValueForID(typeId, _, l) = codeId(typeId, l) val eqCode = if equality then makeEq(t, level, fn (typeId, _, l) => codeId(typeId, l), polyVarMap) else noEquality val boxedCode = boxednessForType(t, level, getTypeValueForID, polyVarMap) val printCode = if printity then printerForType(t, level, polyVarMap) else noPrinter val sizeCode = sizeForType(t, level, getTypeValueForID, polyVarMap) in createTypeValue{ eqCode=eqCode, printCode=printCode, boxedCode=boxedCode, sizeCode=sizeCode} end else (* If we don't require the equality or print function we can use the cache. *) case findCachedTypeCode(polyVarMap, t) of SOME (code, _) => code level | NONE => let val maxCache = getMaxDepth polyVarMap (t, 1) val cacheEntry = List.nth(polyVarMap, List.length polyVarMap - maxCache) val { cache, mkAddr, level=decLevel, ...} = cacheEntry local open TypeValue val boxedCode = boxednessForType(t, decLevel, fn (typeId, _, l) => codeId(typeId, l), polyVarMap) val sizeCode = sizeForType(t, decLevel, fn (typeId, _, l) => codeId(typeId, l), polyVarMap) in val typeValue = createTypeValue{ eqCode=noEquality, printCode=noPrinter, boxedCode=boxedCode, sizeCode=sizeCode} end (* Make a new entry and put it in the cache. *) val decAddr = mkAddr 1 val () = cache := {decCode = mkDec(decAddr, typeValue), typeOf = t, address = decAddr } :: !cache in mkLoad(decAddr, level, decLevel) end in mkEval(code level, List.map makePolyParameter sourceTypes) end (* For now limit this to equality types. *) fun applyToInstance(sourceTypes, level, polyVarMap, code) = applyToInstance'( List.filter(fn {equality, ...} => not justForEqualityTypes orelse equality) sourceTypes, level, polyVarMap, code) structure Sharing = struct type typeId = typeId type codetree = codetree type types = types type typeConstrs= typeConstrs type typeConstrSet=typeConstrSet type typeVarForm=typeVarForm type typeVarMap = typeVarMap type codeBinding = codeBinding type level = level end end; diff --git a/mlsource/MLCompiler/TYPE_TREE.ML b/mlsource/MLCompiler/TYPE_TREE.ML index d5b4460f..e2314b42 100644 --- a/mlsource/MLCompiler/TYPE_TREE.ML +++ b/mlsource/MLCompiler/TYPE_TREE.ML @@ -1,3266 +1,3264 @@ (* Original Poly version: Title: Operations on type structures. Author: Dave Matthews, Cambridge University Computer Laboratory Copyright Cambridge University 1985 ML translation and other changes: Copyright (c) 2000 Cambridge University Technical Services Limited Further development: - Copyright (c) 2000-9, 2012-2018 David C.J. Matthews + Copyright (c) 2000-9, 2012-2018, 2020 David C.J. Matthews This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) functor TYPE_TREE ( structure LEX : LEXSIG structure STRUCTVALS : STRUCTVALSIG; structure PRETTY : PRETTYSIG structure CODETREE : CODETREESIG where type machineWord = Address.machineWord structure EXPORTTREE: EXPORTTREESIG; structure DEBUG: DEBUGSIG structure UTILITIES : sig val mapTable: ('a * 'a -> bool) -> {enter: 'a * 'b -> unit, lookup: 'a -> 'b option} val splitString: string -> { first:string, second:string } end; structure MISC : sig exception InternalError of string; val lookupDefault : ('a -> 'b option) -> ('a -> 'b option) -> 'a -> 'b option end; sharing LEX.Sharing = PRETTY.Sharing = EXPORTTREE.Sharing = STRUCTVALS.Sharing = CODETREE.Sharing ) : TYPETREESIG = (*****************************************************************************) (* TYPETREE functor body *) (*****************************************************************************) struct open MISC; open PRETTY; open STRUCTVALS; open LEX; open UTILITIES; open CODETREE; open EXPORTTREE (* added 16/4/96 SPF *) fun sameTypeVar (TypeVar x, TypeVar y) = sameTv (x, y) | sameTypeVar _ = false; fun isTypeVar (TypeVar _) = true | isTypeVar _ = false; fun isFunctionType (FunctionType _) = true | isFunctionType _ = false; fun isEmpty EmptyType = true | isEmpty _ = false; fun isBadType BadType = true | isBadType _ = false; val emptyType = EmptyType; fun typesTypeVar (TypeVar x) = x | typesTypeVar _ = raise Match; fun typesFunctionType (FunctionType x) = x | typesFunctionType _ = raise Match; (* This is really left over from an old definition. *) fun tcEquivalent(TypeConstrs{identifier = TypeId {idKind = TypeFn(_, result), ...}, ...}) = result | tcEquivalent _ = raise InternalError "tcEquivalent: Not a type function" (* A type construction is the application of a type constructor to a sequence of types to yield a type. A construction may have a nil list if it is a single type identifier such as ``int''. *) (* When a type constructor is encountered in the first pass this entry is put in. Subsequently a type constructor entry will be assigned to it so that the types can be checked. *) (*************) fun mkTypeVar (level, equality, nonunifiable, printable) = TypeVar (makeTv {value=emptyType, level=level, equality=equality, nonunifiable=nonunifiable, printable=printable}); fun mkTypeConstruction (name, typc, args, locations) = TypeConstruction {name = name, constr = typc, args = args, locations = locations} local (* Turn a tuple into a record of the form {1=.., 2=... }*) fun maptoRecord ([], _) = [] | maptoRecord (H::T, i) = {name=Int.toString i, typeof=H} :: maptoRecord (T,i+1) in fun mkProductType (typel: types list) = let val fields = maptoRecord (typel, 1) in LabelledType {recList = fields, fullList = FieldList(List.map #name fields, true)} end end fun mkFunctionType (arg, result) = FunctionType {arg = arg, result = result}; fun mkOverloadSet [constr] = (* If there is just a single constructor in the set we make a type construction from it. *) mkTypeConstruction(tcName constr, constr, nil, []) | mkOverloadSet constrs = let (* Make a type variable and point this at the overload set so we can narrow down the overloading. *) val var = mkTypeVar (generalisable, false, false, false) val set = OverloadSet {typeset=constrs}; in tvSetValue (typesTypeVar var, set); var end fun mkLabelled (l, frozen) = let val final = FieldList(map #name l, frozen) val lab = LabelledType {recList = l, fullList = if frozen then final else FlexibleList(ref final) } in if frozen then lab else let (* Use a type variable so that the record can be expanded. This also provides a model (equality etc). for any fields that are added later. *) val var = mkTypeVar (generalisable, false, false, false) val () = if isTypeVar var then tvSetValue (typesTypeVar var, lab) else (); in var end end (* Must remove leading zeros because the labels are compared by string comparison. *) fun mkLabelEntry (name, t) = let fun stripZeros s = if size s <= 1 orelse String.str(String.sub(s, 0)) <> "0" then s else stripZeros (String.substring(s, 1, size s-1)); in {name = stripZeros name, typeof = t} end; (* Functions to construct the run-time representations of type constructor values, type values and value constructors. These are all tuples and centralising the code here avoids having the offsets as integers at various places. Monotype constructor and type values are almost the same except that type values have the printer entry as the function whereas monotype constructors have the print entry as a ref pointing to the function, allowing addPrettyPrint to set a printer for the type. The entries for polytypes are functions that take the type values as arguments and return the corresponding values. *) structure TypeValue = struct val equalityOffset = 0 and printerOffset = 1 and boxnessOffset = 2 and sizeOffset = 3 local (* Values used to represent boxness. *) val boxedRepNever = 0w1 (* Never boxed, always tagged e.g. bool *) and boxedRepAlways = 0w2 (* Always boxed, never tagged e.g. function types *) and boxedRepEither = 0w3 (* Either boxed or tagged e.g. (arbitrary precision) int *) fun make n = mkConst(Address.toMachineWord n) fun isCode n = - mkInlproc(mkEqualWord(mkLoadArgument 0, make n), 1, "test-box", [], 0) + mkInlproc(mkEqualTaggedWord(mkLoadArgument 0, make n), 1, "test-box", [], 0) in val boxedNever = make boxedRepNever and boxedAlways = make boxedRepAlways and boxedEither = make boxedRepEither (* Test for boxedness. This must be applied to the value extracted from the "boxedness" field after applying to any base type arguments in the case of a polytype constructor. *) val isBoxedNever = isCode boxedRepNever and isBoxedAlways = isCode boxedRepAlways and isBoxedEither = isCode boxedRepEither end (* Sizes are always a single word. *) val singleWord = mkConst(Address.toMachineWord 0w1) fun extractEquality idCode = mkInd(equalityOffset, idCode) and extractPrinter idCode = mkInd(printerOffset, idCode) and extractBoxed idCode = mkInd(boxnessOffset, idCode) and extractSize idCode = mkInd(sizeOffset, idCode) fun createTypeValue{eqCode, printCode, boxedCode, sizeCode} = mkTuple[eqCode, printCode, boxedCode, sizeCode] end (* Value constructors are represented by tuples, either pairs for nullary constructors or triples for constructors with arguments. For nullary functions the "injection" function is actually the value itself. If this is a polytype all the entries are functions that take the type values for the base types as arguments. *) structure ValueConstructor = struct val testerOffset = 0 val injectorOffset = 1 val projectorOffset = 2 fun extractTest constrCode = mkInd(testerOffset, constrCode) and extractInjection constrCode = mkInd(injectorOffset, constrCode) and extractProjection constrCode = mkInd(projectorOffset, constrCode) fun createValueConstr{testMatch, injectValue, projectValue} = mkTuple[testMatch, injectValue, projectValue] fun createNullaryConstr{ testMatch, constrValue } = mkTuple[testMatch, constrValue] end (* Eqtypes with built-in equality functions. The printer functions are all replaced in the basis. *) local open Address PRETTY TypeValue fun defaultMonoTypePrinter _ = PrettyString "?" fun defaultPolyTypePrinter _ _ = PrettyString "?" fun eqAndPrintCode (eqCode, nArgs, boxed) = let val code = if nArgs = 0 then createTypeValue{ eqCode=eqCode, printCode=mkConst (toMachineWord (ref defaultMonoTypePrinter)), boxedCode = boxed, sizeCode = singleWord } else createTypeValue{ eqCode=mkInlproc(eqCode, nArgs, "eq-helper()", [], 0), printCode=mkConst (toMachineWord (ref defaultPolyTypePrinter)), boxedCode = mkInlproc(boxed, nArgs, "boxed-helper()", [], 0), sizeCode = mkInlproc(singleWord, nArgs, "size-helper()", [], 0) } in Global (genCode(code, [], 0) ()) end fun makeConstr(name, fullName, eqFun, boxed) = makeTypeConstructor (name, [], makeFreeId(0, eqAndPrintCode(eqFun, 0, boxed), true, basisDescription fullName), [DeclaredAt inBasis]) (* since code generator relies on these representations, we may as well export them *) - (* wordEq is used both for tagged words and for pointer equality *) - val wordEq = equalWordFn + (* Strings are now always vectors whose first word is the length. The old special case for single-character strings has been removed. *) local val stringEquality = mkInlproc( (* This previously checked for pointer equality first. That has been removed. Test the lengths first and only do the byte comparison if they are the same. This seems to save more time than including the length word in the byte comparison. *) mkCand( - mkEqualWord( + mkEqualPointerOrWord( (* Because we're not actually tagging these we use pointerEq here. *) mkLoadOperation(LoadStoreUntaggedUnsigned, mkLoadArgument 0, CodeZero), mkLoadOperation(LoadStoreUntaggedUnsigned, mkLoadArgument 1, CodeZero)), mkBlockOperation{kind=BlockOpEqualByte, leftBase=mkLoadArgument 0, rightBase=mkLoadArgument 1, leftIndex=mkConst(toMachineWord wordSize), rightIndex=mkConst(toMachineWord wordSize), (* Use argument 1 here rather than 0. We could use either but this works better when we're using equality for pattern matching since it gets the length of the constant string. It also works better for the, to me, more natural ordering of variable=constant. *) length=mkLoadOperation(LoadStoreUntaggedUnsigned, mkLoadArgument 1, CodeZero) } ), 2, "stringEquality", [], 0) in val stringEquality = stringEquality end local (* Arbitrary precision values are normalised so if a value can be represented as a tagged fixed precision value it will be. Unlike strings it is much more likely that the value will be short so we generate equality as a test that handles the short case as inline code and the long case as a function call. If either argument is a short constant this will be optimised away so the test will reduce to a test on whether the value equals the constant. *) val intEquality = mkEnv( [mkDec(0, (* Long-form equality - should not be inlined. *) mkProc( (* Equal if signs are the same ... *) mkCand( - mkEqualWord( + mkEqualTaggedWord( mkUnary(BuiltIns.MemoryCellFlags, mkLoadArgument 0), mkUnary(BuiltIns.MemoryCellFlags, mkLoadArgument 1) ), mkEnv( [mkDec(0, mkUnary(BuiltIns.MemoryCellLength, mkLoadArgument 1))], mkCand( (* ... and the lengths are equal ... *) - mkEqualWord( + mkEqualTaggedWord( mkUnary(BuiltIns.MemoryCellLength, mkLoadArgument 0), mkLoadLocal 0 ), (* ... and they're byte-wise equal .*) mkBlockOperation{kind=BlockOpEqualByte, leftBase=mkLoadArgument 0, rightBase=mkLoadArgument 1, leftIndex=CodeZero, rightIndex=CodeZero, length=mkBinary(BuiltIns.WordArith BuiltIns.ArithMult, mkConst(toMachineWord RunCall.bytesPerWord), mkLoadLocal 0)} ) ) ), 2, "arbitraryPrecisionEquality", [], 1) ) ], mkInlproc( mkCor( (* Either they're equal... *) - (* We can't use mkEqualWord here - because if we have a short constant here the code will reduce to the - simple equality. If we used mkEqualWord it could further reduce to - an indexed case but the code for that does not handle arbitrary - precision values in the lo *) - mkEqualArbShort(mkLoadArgument 0, mkLoadArgument 1), + (* N.B. The values could be short or long. That's particularly important + if we have a series of tests against short constants. If we convert it to + an indexed case we MUST check that the value is short before computing + the index. *) + mkEqualPointerOrWord(mkLoadArgument 0, mkLoadArgument 1), (* .. or if either is short the result is false ... *) mkCand( mkCand( mkNot(mkIsShort(mkLoadArgument 0)), mkNot(mkIsShort(mkLoadArgument 1)) ), (* ... otherwise we have to test the vectors. *) mkEval(mkLoadClosure 0, [mkLoadArgument 0, mkLoadArgument 1]) ) ), 2, "intInfEquality", [mkLoadLocal 0], 0) ) in (* Code-generate the function and return the inline part. We need to set the maximum inline size here to ensure the long form code is not inlined. It would be better to have a way of turning off inlining for specific functions. *) val intEquality = genCode(intEquality, [Universal.tagInject DEBUG.maxInlineSizeTag 5], 1) () end in - val fixedIntConstr = makeConstr("int", "FixedInt.int", wordEq, boxedNever) (* Fixed precision is always short *) + val fixedIntConstr = makeConstr("int", "FixedInt.int", equalTaggedWordFn, boxedNever) (* Fixed precision is always short *) val intInfConstr = makeConstr("int", "IntInf.int", intEquality, boxedEither) - val charConstr = makeConstr("char", "char", wordEq, boxedNever) (* Always short *) + val charConstr = makeConstr("char", "char", equalTaggedWordFn, boxedNever) (* Always short *) val stringConstr = makeConstr("string", "string", stringEquality, boxedEither (* Single chars are unboxed. *)) - val wordConstr = makeConstr("word", "word", wordEq, boxedNever) + val wordConstr = makeConstr("word", "word", equalTaggedWordFn, boxedNever) (* Ref is a datatype with a single constructor. The constructor is added in INITIALISE. Equality is special for "'a ref", "'a array" and "'a Array2.array". They permit equality even if the 'a is not an eqType. *) val refConstr = makeTypeConstructor ("ref", [makeTv {value=EmptyType, level=generalisable, equality=false, nonunifiable=false, printable=false}], - makeFreeId(1, eqAndPrintCode(wordEq, 1, boxedAlways), true, basisDescription "ref"), + makeFreeId(1, eqAndPrintCode(equalPointerOrWordFn, 1, boxedAlways), true, basisDescription "ref"), [DeclaredAt inBasis]); val arrayConstr = makeTypeConstructor ("array", [makeTv {value=EmptyType, level=generalisable, equality=false, nonunifiable=false, printable=false}], - makeFreeId(1, eqAndPrintCode(wordEq, 1, boxedAlways), true, basisDescription "Array.array"), + makeFreeId(1, eqAndPrintCode(equalPointerOrWordFn, 1, boxedAlways), true, basisDescription "Array.array"), [DeclaredAt inBasis]); val array2Constr = makeTypeConstructor ("array", [makeTv {value=EmptyType, level=generalisable, equality=false, nonunifiable=false, printable=false}], - makeFreeId(1, eqAndPrintCode(wordEq, 1, boxedAlways), true, basisDescription "Array2.array"), + makeFreeId(1, eqAndPrintCode(equalPointerOrWordFn, 1, boxedAlways), true, basisDescription "Array2.array"), [DeclaredAt inBasis]); val byteArrayConstr = makeTypeConstructor ("byteArray", [], - makeFreeId(0, eqAndPrintCode(wordEq, 0, boxedAlways), true, basisDescription "byteArray"), + makeFreeId(0, eqAndPrintCode(equalPointerOrWordFn, 0, boxedAlways), true, basisDescription "byteArray"), [DeclaredAt inBasis]); (* Bool is a datatype. The constructors are added in INITIALISE. *) val boolConstr = makeTypeConstructor - ("bool", [], makeFreeId(0, eqAndPrintCode(wordEq, 0, boxedNever), true, basisDescription "bool"), + ("bool", [], makeFreeId(0, eqAndPrintCode(equalTaggedWordFn, 0, boxedNever), true, basisDescription "bool"), [DeclaredAt inBasis]); end (* These polytypes allow equality even if the type argument is not an equality type. *) fun isPointerEqType id = sameTypeId (id, tcIdentifier refConstr) orelse sameTypeId (id, tcIdentifier arrayConstr) orelse sameTypeId (id, tcIdentifier array2Constr) orelse sameTypeId (id, tcIdentifier byteArrayConstr) (* Non-eqtypes *) local open Address PRETTY TypeValue fun makeType(name, descr, boxed) = let fun defaultPrinter _ = PrettyString "?" val code = createTypeValue{ eqCode=CodeZero (* No equality. *), printCode=mkConst (toMachineWord (ref defaultPrinter)), boxedCode=boxed, sizeCode=singleWord } in makeTypeConstructor ( name, [], makeFreeId(0, Global (genCode(code, [], 0) ()), false, descr), [DeclaredAt inBasis]) end in val realConstr = makeType("real", basisDescription "real", boxedAlways(* Currently*)) (* Not an eqtype in ML97. *) (* Short real: Real32.real *) val floatConstr = makeType("real", basisDescription "real", if RunCall.bytesPerWord <= 0w4 then boxedAlways else boxedNever) val exnConstr = makeType("exn", basisDescription "exn", boxedAlways); (* "undefConstr" is used as a place-holder during parsing for the actual type constructor. If the type constructor is not found this may appear in an error message. *) val undefConstr = makeType("undefined", { location = inBasis, description = "Undefined", name = "undefined" }, boxedEither); end (* The unit type is equivalent to the empty record. *) val unitConstr = makeTypeConstructor ("unit", [], makeTypeFunction({ location = inBasis, description = "unit", name = "unit" }, ([], LabelledType {recList = [], fullList = FieldList([], true)})), [DeclaredAt inBasis]); (* Type identifiers bound to standard type constructors. *) val unitType = mkTypeConstruction ("unit", unitConstr, [], []) val fixedIntType = mkTypeConstruction ("int", fixedIntConstr, [], []) val stringType = mkTypeConstruction ("string", stringConstr, [], []) val boolType = mkTypeConstruction ("bool", boolConstr, [], []) val exnType = mkTypeConstruction ("exn", exnConstr, [], []) fun isUndefined cons = sameTypeId (tcIdentifier cons, tcIdentifier undefConstr); val isUndefinedTypeConstr = isUndefined (* Test if a type is the undefined constructor. *) fun isUndefinedType(TypeConstruction{constr, ...}) = isUndefined constr | isUndefinedType _ = false (* Similar to alphabetic ordering except that shorter labels come before longer ones. This has the advantage that numerical labels are compared by their numerical order i.e. 1 < 2 < 10 whereas alphabetic ordering puts "1" < "10" < "2". *) fun compareLabels (a : string, b : string) : int = if size a = size b then if a = b then 0 else if a < b then ~1 else 1 else if size a < size b then ~1 else 1; (* Sort using the label ordering. A simple sort routine - particularly if the list is already sorted. *) fun sortLabels [] = [] | sortLabels (s::rest) = let fun enter s _ [] = [s] | enter s name (l as ( (h as {name=hname, ...}) :: t)) = let val comp = compareLabels (name, hname); in if comp <= 0 then s :: l else h :: enter s name t end; in enter s (#name s) (sortLabels rest) end (* Chains down a list of type variables returning the type they are bound to. As a side-effect it also points all the type variables at this type to reduce the need for future chaining and to free unused type variables. Normally a type variable points to at most one other, which then points to "empty". However if we have unified two type variables by pointing one at the other, there may be type variables which pointed to the first and which cannot be found and redirected at the second until they themselves are examined. *) fun eventual (t as (TypeVar tv)) : types = let (* Note - don't change the level/copy information - the only type variable with this correct is the one at the end of the list. *) val oldVal = tvValue tv val newVal = eventual oldVal; (* Search that *) in (* Update the type variable to point to the last in the chain. We don't do this if the value hasn't changed. The reason for that was that assignment to refs in the database in the old persistent store system was very expensive and we wanted to avoid unnecessary assignments. This special case could probably be removed. *) if PolyML.pointerEq(oldVal, newVal) then () else tvSetValue (tv, newVal); (* Put it on *) case newVal of EmptyType => t (* Not bound to anything - return the type variable *) | LabelledType (r as { recList, fullList }) => if List.length recList = List.length(recordFields r) then (* All the generic fields are present so we don't need to do anything. *) if recordIsFrozen r then newVal else t else (* We need to add fields from the generic. *) let (* Add any fields from the generic that aren't present in this instance. *) fun createNewField name = { name = name, (* The new type variable has to be created with the same properties as if we had first generalised it from the generic and then unified with this instance. The level is inherited from the instance since the generic will always have level = generalisable. Nonunifiable must be false. *) typeof = mkTypeVar (tvLevel tv, tvEquality tv, false, tvPrintity tv)} fun addToInstance([], []) = [] | addToInstance(generic :: geRest, []) = createNewField generic :: addToInstance(geRest, []) | addToInstance([], instance) = instance (* This case can occur if we are producing an error message because of a type-incorrect program so we just ignore it. *) | addToInstance(generic :: geRest, inst as instance :: iRest) = let val order = compareLabels (generic, #name instance); in if order = 0 (* Equal *) then instance :: addToInstance(geRest, iRest) else if order < 0 (* generic name < instance name *) then createNewField generic :: addToInstance(geRest, inst) else (* This is another case that can occur with type-incorrect code. *) instance :: addToInstance(generic :: geRest, iRest) end val newList = addToInstance(recordFields r, recList) val newRecord = LabelledType {recList = newList, fullList = fullList} in tvSetValue(tv, newRecord); if recordIsFrozen r then newRecord else t end | OverloadSet _ => t (* Return the set of types. *) | _ => newVal (* Return the type it is bound to *) end | eventual t (* not a type variable *) = t; (* Apply a function to every element of a type. *) fun foldType f = let fun foldT typ v = let val t = eventual typ; val res = f t v; (* Process this entry. *) in case t of TypeVar tv => foldT (tvValue tv) res | TypeConstruction {args, ...} => (* Then process the arguments. *) List.foldr (fn (t, v) => foldT t v) res args | FunctionType {arg, result} => foldT arg (foldT result res) | LabelledType {recList,...} => List.foldr (fn ({ typeof, ... }, v) => foldT typeof v) res recList | BadType => res | EmptyType => res | OverloadSet _ => res end in foldT end; (* Checks to see whether a labelled record is in the form of a product i.e. 1=, 2= We only need this for prettyprinting. Zero-length records (i.e. unit) and singleton records are not considered as tuples. *) fun isProductType(LabelledType(r as {recList=recList as _::_::_, ...})) = let fun isRec [] _ = true | isRec ({name, ...} :: l) n = name = Int.toString n andalso isRec l (n+1) in recordIsFrozen r andalso isRec recList 1 end | isProductType _ = false; (* Test to see is a type constructor is in an overload set. *) fun isInSet(tcons: typeConstrs, (H::T): typeConstrs list) = sameTypeId (tcIdentifier tcons, tcIdentifier H) orelse isInSet(tcons, T) | isInSet(_, []: typeConstrs list) = false val prefInt = ref fixedIntConstr (* Returns the preferred overload if there is one. *) fun preferredOverload typeset = if isInSet(!prefInt, typeset) then SOME(!prefInt) else if isInSet(realConstr, typeset) then SOME realConstr else if isInSet(wordConstr, typeset) then SOME wordConstr else if isInSet(charConstr, typeset) then SOME charConstr else if isInSet(stringConstr, typeset) then SOME stringConstr else NONE fun setPreferredInt c = prefInt := c fun equalTypeIds(x, y) = let (* True if two types are equal. *) fun equalTypes (TypeConstruction{constr=xVal, args=xArgs, ...}, TypeConstruction{constr=yVal, args=yArgs, ...}) = equalTypeIds(tcIdentifier xVal, tcIdentifier yVal) andalso equalTypeLists (xArgs, yArgs) | equalTypes (FunctionType x, FunctionType y) = equalTypes (#arg x, #arg y) andalso equalTypes (#result x, #result y) | equalTypes (LabelledType x, LabelledType y) = recordIsFrozen x andalso recordIsFrozen y andalso equalRecordLists (#recList x, #recList y) | equalTypes (TypeVar x, TypeVar y) = sameTv (x, y) | equalTypes (EmptyType, EmptyType) = true | equalTypes _ = false and equalTypeLists ([], []) = true | equalTypeLists (x::xs, y::ys) = equalTypes(x, y) andalso equalTypeLists (xs, ys) | equalTypeLists _ = false and equalRecordLists ([], []) = true | equalRecordLists (x::xs, y::ys) = #name x = #name y andalso equalTypes(#typeof x, #typeof y) andalso equalRecordLists (xs, ys) | equalRecordLists _ = false in case (x, y) of (TypeId{idKind=TypeFn(_, xEquiv), ...}, TypeId{idKind=TypeFn(_, yEquiv), ...}) => equalTypes(xEquiv, yEquiv) | _ => sameTypeId(x, y) end (* See if the types are the same. This is a bit of a fudge, but saves carrying around a flag saying whether the structures were copied. This is only an optimisation. If the values are different it will not go wrong. *) val identical : types * types -> bool = PolyML.pointerEq and identicalConstr : typeConstrs * typeConstrs -> bool = PolyML.pointerEq and identicalList : 'a list * 'a list -> bool = PolyML.pointerEq (* Copy a type, avoiding copying type structures unnecessarily. Used to make new type variables for all distinct type variables when generalising polymorphic functions, and to make new type stamps for type constructors when generalising signatures. *) fun copyType (at, copyTypeVar, copyTypeConstr) = let fun copyList [] = [] | copyList (l as (h :: t)) = let val h' = copyType (h, copyTypeVar, copyTypeConstr); val t' = copyList t; in if identical (h', h) andalso identicalList (t', t) then l else h' :: t' end (* copyList *); fun copyRecordList [] = [] | copyRecordList (l as ({name, typeof} :: t)) = let val typeof' = copyType (typeof, copyTypeVar, copyTypeConstr); val t' = copyRecordList t; in if identical (typeof', typeof) andalso identicalList (t', t) then l else {name=name, typeof=typeof'} :: t' end (* copyList *); val atyp = eventual at; in case atyp of TypeVar _ => (* Unbound type variable, flexible record or overloading. *) copyTypeVar atyp | TypeConstruction {constr, args, locations, ...} => let val copiedArgs = copyList args; val copiedConstr = copyTypeConstr constr; (* Use the name from the copied constructor. This will normally be the same as the original EXCEPT in the case where we are using copyType to generate copies of the value constructors of replicated datatypes. *) val copiedName = tcName copiedConstr in if identicalList (copiedArgs, args) andalso identicalConstr (copiedConstr, constr) then atyp else (* Must copy it. *) mkTypeConstruction (copiedName, copiedConstr, copiedArgs, locations) end | FunctionType {arg, result} => let val copiedArg = copyType (arg, copyTypeVar, copyTypeConstr); val copiedRes = copyType (result, copyTypeVar, copyTypeConstr); in if identical (copiedArg, arg) andalso identical (copiedRes, result) then atyp else FunctionType {arg = copiedArg, result = copiedRes} end | LabelledType {recList, fullList} => (* Rigid labelled records only. Flexible ones are treated as type vars. *) let val copiedList = copyRecordList recList in if identicalList (copiedList, recList) then atyp else LabelledType {recList = copiedList, fullList = fullList} end | EmptyType => EmptyType | BadType => BadType | OverloadSet _ => raise InternalError "copyType: OverloadSet found" end (* copyType *); (* Copy a type constructor if it is Bound and in the required range. If this refers to a type function copies that as well. Does not copy value constructors. *) fun copyTypeConstrWithCache (tcon, typeMap, _, mungeName, cache) = case tcIdentifier tcon of TypeId{idKind = TypeFn(args, equiv), description, access, ...} => let val copiedEquiv = copyType(equiv, fn x => x, fn tcon => copyTypeConstrWithCache (tcon, typeMap, fn x => x, mungeName, cache)) in if identical (equiv, copiedEquiv) then tcon (* Type is identical and we don't want to change the name. *) else (* How do we find a type function? *) makeTypeConstructor (mungeName(tcName tcon), args, TypeId { access = access, description = description, idKind = TypeFn(args, copiedEquiv)}, tcLocations tcon) end | id => ( case typeMap id of NONE => ( (*print(concat[tcName tcon, " not copied\n"]);*) tcon (* No change *) ) | SOME newId => let val name = #second(splitString (tcName tcon)) (* We must only match here if they're really the same. *) fun cacheMatch tc = equalTypeIds(tcIdentifier tc, newId) andalso #second(splitString(tcName tc)) = name in case List.find cacheMatch cache of SOME tc => ( (*print(concat[tcName tcon, " copied as ", tcName tc, "\n"]);*) tc (* Use the entry from the cache. *) ) | NONE => (* Either a hidden identifier or alternatively this can happen as part of the matching process. When matching a structure to a signature we first match up the type constructors then copy the type of each value replacing bound type IDs with the actual IDs as part of the checking process. We will return SOME newId but we don't have a cache so return NONE for List.find. *) let val newName = mungeName(tcName tcon) in (*print(concat[tcName tcon, " not cached\n"]);*) makeTypeConstructor(newName, tcTypeVars tcon, newId, tcLocations tcon) end end ) (* Exported version. *) fun copyTypeConstr (tcon, typeMap, copyTypeVar, mungeName) = copyTypeConstrWithCache(tcon, typeMap, copyTypeVar, mungeName, []) (* Compose typeID maps. If the first map returns a Bound id we apply the second otherwise just return the result of the first. *) fun composeMaps(m1, m2) n = let fun map2 (TypeId{idKind=Bound{ offset, ...}, ...}) = m2 offset | map2 (id as TypeId{idKind=Free _, ...}) = id | map2 (TypeId{idKind=TypeFn(args, equiv), access, description, ...}) = let fun copyId(TypeId{idKind=Free _, ...}) = NONE | copyId id = SOME(map2 id) (* If it's a type function e.g. this was a "where type" we have to apply the map to any type identifiers in the type. *) val copiedEquiv = copyType(equiv, fn x => x, fn tcon => copyTypeConstr (tcon, copyId, fn x => x, fn y => y)) in TypeId{idKind = TypeFn(args, copiedEquiv), access=access, description=description} end in map2(m1 n) end (* Basic procedure to print a type structure. *) type printTypeEnv = { lookupType: string -> (typeConstrSet * (int->typeId) option) option, lookupStruct: string -> (structVals * (int->typeId) option) option} val emptyTypeEnv = { lookupType = fn _ => NONE, lookupStruct = fn _ => NONE } (* Test whether two type constructors are the same after mapping. This is used to try to find the correct "path" to a type constructor when printing. *) fun eqTypeConstrs(xTypeCons, xMap, yTypeCons, yMap) = let fun id x = x fun copyId (SOME mapTypeId) (TypeId{idKind=Bound{ offset, ...}, ...}) = SOME(mapTypeId offset) | copyId _ _ = NONE val mappedX = copyTypeConstr(xTypeCons, copyId xMap, id, id) and mappedY = copyTypeConstr(yTypeCons, copyId yMap, id, id) in equalTypeIds(tcIdentifier mappedX, tcIdentifier mappedY) end (* prints a block of items *) fun tDisp (t : types, depth : FixedInt.int, typeVarName : typeVarForm -> string, env: printTypeEnv, sigMap: (int->typeId)option) : pretty = let (* prints a block of items *) fun dispP (t : types, depth : FixedInt.int) : pretty = let (* prints a block of items *) fun parenthesise depth t = if depth <= 1 then PrettyString "..." else PrettyBlock (0, false, [], [ PrettyString "(", dispP (t, depth - 1), PrettyString ")" ]); (* prints a sequence of items *) fun prettyList [] _ _: pretty list = [] | prettyList [H] depth separator = let val v = eventual H; in if separator = "*" andalso (isFunctionType v orelse isProductType v) then (* Must bracket the expression *) [parenthesise depth v] else [dispP (v, depth)] end | prettyList (H :: T) depth separator = if depth <= 0 then [PrettyString "..."] else let val v = eventual H; in PrettyBlock (0, false, [], [(if separator = "*" andalso (isFunctionType v orelse isProductType v) then (* Must bracket the expression *) parenthesise depth v else dispP (v, depth)), PrettyBreak (if separator = "," then 0 else 1, 0), PrettyString separator ]) :: PrettyBreak (1, 0) :: prettyList T (depth - 1) separator end; val typ = eventual t; (* Find the real type structure *) in case typ of TypeVar tyVar => let val tyVal : types = tvValue tyVar; in case tyVal of EmptyType => PrettyString (typeVarName tyVar) | _ => dispP (tyVal, depth) end (* Type construction. *) | TypeConstruction {args, name, constr=typeConstructor, ...} => let val constrName = (* Use the type constructor name unless we're had an error. *) if isUndefined typeConstructor then name else tcName typeConstructor (* There are three possible cases: we may not find any type with the name, we may look up the name and find the type or we may look up the name and find a different type. *) datatype isFound = NotFound | FoundMatch | FoundNotMatch (* If we're printing a value that refers to a type constructor we want to print the correct amount of any structure prefix for the current context. *) fun findType (_, []) = NotFound | findType ({ lookupType, ... }, [typeName]) = ( (* This must be the name of a type. *) case lookupType typeName of SOME (t, map) => if eqTypeConstrs(typeConstructor, sigMap, tsConstr t, map) then FoundMatch else FoundNotMatch | NONE => NotFound ) | findType ({ lookupStruct, ... }, structName :: tail) = ( (* This must be the name of a structure. Does it contain our type? *) case lookupStruct structName of SOME(Struct { signat, ...}, map) => let val Signatures { tab, typeIdMap, ...} = signat val Env { lookupType, lookupStruct, ...} = makeEnv tab val newMap = case map of SOME map => composeMaps(typeIdMap, map) | NONE => typeIdMap fun subLookupType s = case lookupType s of NONE => NONE | SOME t => SOME(t, SOME newMap) fun subLookupStruct s = case lookupStruct s of NONE => NONE | SOME t => SOME(t, SOME newMap) in findType({lookupType=subLookupType, lookupStruct=subLookupStruct}, tail) end | NONE => NotFound ) (* See if we have this type in the current environment or in some structure in the current environment. The name we have may be a full structure path. *) fun nameToList ("", l) = (l, NotFound) (* Not there. *) | nameToList (s, l) = let val { first, second } = splitString s val currentList = second :: l in case findType(env, currentList) of FoundMatch => (currentList, FoundMatch) | FoundNotMatch => ( case nameToList(first, currentList) of result as (_, FoundMatch) => result | (l, _) => (l, FoundNotMatch) ) | NotFound => nameToList(first, currentList) end (* Try the type constructor name first. This is usually accurate. If not fall back to the type identifier. This may be needed in rarer cases. *) val names = case nameToList(constrName, []) of (names, FoundMatch) => names (* Found the type constructor name. *) | (names, f) => let (* Try the type identifier name. *) val TypeId { description = { name=idName, ...}, ...} = case (sigMap, tcIdentifier typeConstructor) of (SOME map, TypeId{idKind=Bound{offset, ...}, ...}) => map offset | (_, id) => id (* Only add "?" if we actually found a type with the required name but it wasn't the right one. This allows us to print a sensible result where the type has been shadowed but doesn't affect situations such as where we create a unique type name for a free type variable. *) fun addQuery n = case f of FoundNotMatch => "?" :: n | _ => n in if idName = "" then addQuery names else case nameToList(idName, []) of (idNames, FoundMatch) => idNames | (_, _) => addQuery names (* Print it as "?.t". This isn't ideal but will help in situations where we have redefined "t". *) end val newName = String.concatWith "." names (* Get the declaration position for the type constructor. *) val constrContext = if isUndefined typeConstructor then [] else ( case List.find(fn DeclaredAt _ => true | _ => false) (tcLocations typeConstructor) of SOME(DeclaredAt loc) => [ContextLocation loc] | _ => [] ) val constructorEntry = PrettyBlock(0, false, constrContext, [PrettyString newName(*constrName*)]) in case args of [] => constructorEntry | args as hd :: tl => let val argVal = eventual hd; in PrettyBlock (0, false, [], [ (* If we have just a single argument and it's just a type constructor or a construction we don't need to parenthesise it. *) if null tl andalso not (isProductType argVal orelse isFunctionType argVal) then dispP (argVal, depth - 1) else if depth <= 1 then PrettyString "..." else PrettyBlock(0, false, [], [PrettyString "(", PrettyBreak (0, 0)] @ prettyList args (depth - 1) "," @ [PrettyBreak (0, 0), PrettyString ")"] ), PrettyBreak(1, 0), constructorEntry (* The constructor. *) ]) end end | FunctionType {arg, result} => if depth <= 0 then PrettyString "..." else (* print out in infix notation *) let val evArg = eventual arg; in PrettyBlock (0, false, [], [ (* If the argument is a function it must be printed as (a-> b)->.. *) if isFunctionType evArg then parenthesise depth evArg else dispP (evArg, depth - 1), PrettyBreak(1, 2), PrettyString "->", PrettyBreak (1, 2), dispP (result, depth - 1) ]) end | LabelledType (r as {recList, ...}) => if depth <= 0 then PrettyString "..." else if isProductType typ then (* Print as a product *) PrettyBlock (0, false, [], (* Print them as t1 * t2 * t3 .... *) prettyList (map (fn {typeof, ...} => typeof) recList) depth "*") else (* Print as a record *) let (* The ordering on fields is designed to allow mixing of tuples and records (e.g. #1). It puts shorter names before longer so that #11 comes after #2 and before #100. For named records it does not make for easy reading so we sort those alphabetically when printing. *) val sortedRecList = Misc.quickSort(fn {name = a, ...} => fn {name = b, ...} => a <= b) recList in PrettyBlock (2, false, [], PrettyString "{" :: (let fun pRec [] _ = [] | pRec ({name, typeof} :: T) depth = if depth <= 0 then [PrettyString "..."] else [ PrettyBlock(0, false, [], [ PrettyBlock(0, false, [], [ PrettyString (name ^ ":"), PrettyBreak(1, 0), dispP(typeof, depth - 1) ] @ (if null T then [] else [PrettyBreak (0, 0), PrettyString ","]) ) ]@ (if null T then [] else PrettyBreak (1, 0) :: pRec T (depth-1)) ) ] in pRec sortedRecList (depth - 1) end) @ [ PrettyString (if recordIsFrozen r then "}" else case recList of [] => "...}" | _ => ", ...}")] ) end | OverloadSet {typeset = []} => PrettyString "no type" | OverloadSet {typeset = tconslist} => (* This typically arises when printing error messages in the second pass because the third pass will select a single type e.g. int where possible. To simplify the messages select a single type if possible. *) ( case preferredOverload tconslist of SOME tcons => dispP(mkTypeConstruction (tcName tcons, tcons,[], []), depth) | NONE => (* Just print the type constructors separated by / *) let fun constrLocation tcons = case List.find(fn DeclaredAt _ => true | _ => false) (tcLocations tcons) of SOME(DeclaredAt loc) => [ContextLocation loc] | _ => [] (* Type constructor with context. *) fun tconsItem tcons = PrettyBlock(0, false, constrLocation tcons, [PrettyString(tcName tcons)]) fun printTCcons [] = [] | printTCcons [tcons] = [tconsItem tcons] | printTCcons (tcons::rest) = tconsItem tcons :: PrettyBreak (0, 0) :: PrettyString "/" :: printTCcons rest in PrettyBlock (0, false, [], printTCcons tconslist) end ) | EmptyType => PrettyString "no type" | BadType => PrettyString "bad" end (* dispP *) in dispP (t, depth) end (* tDisp *); (* Generate unique type-variable names. *) fun varNameSequence () : typeVarForm -> string = (* We need to ensure that every distinct type variable has a distinct name. Each new type variable is given a name starting at "'a" and going on through the alphabet. *) let datatype names = Names of {name: string, entry: typeVarForm} val nameNum = ref ~1 val gNameList = ref [] (* List of names *) in (* If the type is already there return the name we have given it otherwise make a new name and put it in the list. *) fn var => case List.find (fn (Names {entry,...}) => sameTv (entry, var)) (!gNameList) of NONE => (* Not on the list - make a new name *) let fun name num = (if num >= 26 then name (num div 26 - 1) else "") ^ String.str (Char.chr (num mod 26 + Char.ord #"a")) val () = nameNum := !nameNum + 1 val n = (if tvEquality var then "''" else "'") ^ name(!nameNum) (* Should explicit type variables be distinguished? *) in gNameList := Names{name=n, entry=var} :: !gNameList; n end | SOME (Names {name,...}) => name end (* varNameSequence *) (* Print a type (as a block of items) *) fun displayWithMap (t : types, depth : FixedInt.int, env, sigMap) = tDisp (t, depth, varNameSequence (), env, sigMap) and display (t : types, depth : FixedInt.int, env) = tDisp (t, depth, varNameSequence (), env, NONE) (* Print out zero, one or more type variables (unblocked) *) fun printTypeVars([], _, _) = [] (* No type vars i.e. monotype *) | printTypeVars([oneVar], depth, typeV) = (* Single type var. *) [ tDisp (TypeVar oneVar, depth, typeV, emptyTypeEnv, NONE), PrettyBreak (1, 0) ] | printTypeVars(vars, depth, typeV) = (* Must parenthesise them. *) if depth <= 1 then [PrettyString "..."] else [ PrettyBlock(0, false, [], PrettyString "(" :: PrettyBreak(0, 0) :: (let fun pVars vars depth: pretty list = if depth <= 0 then [PrettyString "..."] else if null vars then [] else [ tDisp (TypeVar(hd vars), depth, typeV, emptyTypeEnv, NONE), PrettyBreak (0, 0) ] @ (if null (tl vars) then [] else PrettyString "," :: PrettyBreak (1, 0) :: pVars (tl vars) (depth - 1) ) in pVars vars depth end) @ [PrettyString ")"] ), PrettyBreak (1, 0) ] (* Version used in parsetree. *) fun displayTypeVariables (vars : typeVarForm list, depth : FixedInt.int) = printTypeVars (vars, depth, varNameSequence ()) (* Parse tree for types. This is used to represent types in the source. *) datatype typeParsetree = ParseTypeConstruction of { name: string, args: typeParsetree list, location: location, nameLoc: location, argLoc: location, (* foundConstructor is set to the constructor when it has been looked up. This allows us to get the location where it was declared if we export the parse-tree. *) foundConstructor: typeConstrs ref } | ParseTypeProduct of { fields: typeParsetree list, location: location } | ParseTypeFunction of { argType: typeParsetree, resultType: typeParsetree, location: location } | ParseTypeLabelled of { fields: ((string * location) * typeParsetree * location) list, frozen: bool, location: location } | ParseTypeId of { types: typeVarForm, location: location } | ParseTypeBad (* Place holder for errors. *) fun typeFromTypeParse( ParseTypeConstruction{ args, name, location, foundConstructor = ref constr, ...}) = let val argTypes = List.map typeFromTypeParse args in TypeConstruction {name = name, constr = constr, args = argTypes, locations = [DeclaredAt location]} end | typeFromTypeParse(ParseTypeProduct{ fields, ...}) = mkProductType(List.map typeFromTypeParse fields) | typeFromTypeParse(ParseTypeFunction{ argType, resultType, ...}) = mkFunctionType(typeFromTypeParse argType, typeFromTypeParse resultType) | typeFromTypeParse(ParseTypeLabelled{ fields, frozen, ...}) = let fun makeField((name, _), t, _) = mkLabelEntry(name, typeFromTypeParse t) in mkLabelled(sortLabels(List.map makeField fields), frozen) end | typeFromTypeParse(ParseTypeId{ types, ...}) = TypeVar types | typeFromTypeParse(ParseTypeBad) = BadType fun makeParseTypeConstruction((constrName, nameLoc), (args, argLoc), location) = ParseTypeConstruction{ name = constrName, nameLoc = nameLoc, args = args, argLoc = argLoc, location = location, foundConstructor = ref undefConstr } fun makeParseTypeProduct(recList, location) = ParseTypeProduct{ fields = recList, location = location } fun makeParseTypeFunction(arg, result, location) = ParseTypeFunction{ argType = arg, resultType = result, location = location } fun makeParseTypeLabelled(recList, frozen, location) = ParseTypeLabelled{ fields = recList, frozen = frozen, location = location } fun makeParseTypeId(types, location) = ParseTypeId{ types = types, location = location } fun unitTree location = ParseTypeLabelled{ fields = [], frozen = true, location = location } (* Build an export tree from the parse tree. *) fun typeExportTree(navigation, p: typeParsetree) = let val typeof = typeFromTypeParse p (* Common properties for navigation and printing. *) val commonProps = PTprint(fn d => display(typeof, d, emptyTypeEnv)) :: PTtype typeof :: exportNavigationProps navigation fun asParent () = typeExportTree(navigation, p) in case p of ParseTypeConstruction{ location, nameLoc, args, argLoc, ...} => let (* If the constructor has been bound return the declaration location. We have to attach the declaration location in the right place if this is a polytype e.g. if we have "int list" here we will have the location for "list" which is the second item not the first. *) val (name, decLoc) = case typeof of TypeConstruction { constr, name, ...} => if isUndefined constr then (name, []) else (name, mapLocationProps(tcLocations constr)) | _ => ("", []) (* Error? *) val navNameAndArgs = (* Separate cases for nullary, unary and higher type constructions. *) case args of [] => decLoc (* Singleton e.g. int *) | [oneArg] => let (* Single arg e.g. int list. *) (* Navigate between the type constructor and the argument. Since the arguments come before the constructor we go there first. *) fun getArg () = typeExportTree({parent=SOME asParent, previous=NONE, next=SOME getName}, oneArg) and getName () = getStringAsTree({parent=SOME asParent, previous=SOME getArg, next=NONE}, name, nameLoc, decLoc) in [PTfirstChild getArg] end | args => let (* Multiple arguments e.g. (int, string) pair *) fun getArgs () = (argLoc, exportList(typeExportTree, SOME getArgs) args @ exportNavigationProps{parent=SOME asParent, previous=NONE, next=SOME getName}) and getName () = getStringAsTree({parent=SOME asParent, previous=SOME getArgs, next=NONE}, name, nameLoc, decLoc) in [PTfirstChild getArgs] end in (location, navNameAndArgs @ commonProps) end | ParseTypeProduct{ location, fields, ...} => (location, exportList(typeExportTree, SOME asParent) fields @ commonProps) | ParseTypeFunction{ location, argType, resultType, ...} => (location, exportList(typeExportTree, SOME asParent) [argType, resultType] @ commonProps) | ParseTypeLabelled{ location, fields, ...} => let fun exportField(navigation, label as ((name, nameLoc), t, fullLoc)) = let (* The first position is the label, the second the type *) fun asParent () = exportField (navigation, label) fun getLab () = getStringAsTree({parent=SOME asParent, next=SOME getType, previous=NONE}, name, nameLoc, [PTtype(typeFromTypeParse t)]) and getType () = typeExportTree({parent=SOME asParent, previous=SOME getLab, next=NONE}, t) in (fullLoc, PTfirstChild getLab :: exportNavigationProps navigation) end in (location, exportList(exportField, SOME asParent) fields @ commonProps) end | ParseTypeId{ location, ...} => (location, commonProps) | ParseTypeBad => (nullLocation, commonProps) end fun displayTypeParse(types, depth, env) = display(typeFromTypeParse types, depth, env) (* Associates type constructors from the environment with type identifiers (NOT type variables) *) fun assignTypes (tp : typeParsetree, lookupType : string * location -> typeConstrSet, lex : lexan) = let fun typeFromTypeParse(ParseTypeConstruction{ args, name, location, foundConstructor, ...}) = let (* Assign constructor, then the parameters. *) val TypeConstrSet(constructor, _) = lookupType (name, location) val () = (* Check that it has the correct arity. *) if not (isUndefined constructor) then let val arity : int = tcArity constructor; val num : int = length args; in if arity <> num then (* Give an error message *) errorMessage (lex, location, String.concat["Type constructor (", tcName constructor, ") requires ", Int.toString arity, " type(s) not ", Int.toString num]) else foundConstructor := constructor end else () val argTypes = List.map typeFromTypeParse args in TypeConstruction {name = name, constr = constructor, args = argTypes, locations = [DeclaredAt location]} end | typeFromTypeParse(ParseTypeProduct{ fields, ...}) = mkProductType(List.map typeFromTypeParse fields) | typeFromTypeParse(ParseTypeFunction{ argType, resultType, ...}) = mkFunctionType(typeFromTypeParse argType, typeFromTypeParse resultType) | typeFromTypeParse(ParseTypeLabelled{ fields, frozen, ...}) = let fun makeField((name, _), t, _) = mkLabelEntry(name, typeFromTypeParse t) in mkLabelled(sortLabels(List.map makeField fields), frozen) end | typeFromTypeParse(ParseTypeId{ types, ...}) = TypeVar types | typeFromTypeParse(ParseTypeBad) = BadType in typeFromTypeParse tp end; (* When we have finished processing a list of patterns we need to check that the record is now frozen. *) fun recordNotFrozen (TypeVar t) : bool = (* Follow the chain *) recordNotFrozen (tvValue t) | recordNotFrozen (LabelledType r) = not(recordIsFrozen r) | recordNotFrozen _ = false (* record or type alias *); datatype generalMatch = Matched of {old: typeVarForm, new: types}; fun generaliseTypes (atyp : types, checkTv: typeVarForm->types option) = let val madeList = ref [] (* List of tyVars. *); fun tvs atyp = let val tyVar = typesTypeVar atyp; in case List.find(fn Matched{old, ...} => sameTv (old, tyVar)) (!madeList) of SOME(Matched{new, ...}) => new | NONE => ( case checkTv tyVar of SOME found => found | NONE => let (* Not on the list - make a new name *) (* Make a unifiable type variable even if the original is nonunifiable. *) val n : types = mkTypeVar (generalisable, tvEquality tyVar, false, tvPrintity tyVar) in (* Set the new variable to have the same value as the existing. That is only really needed if we have an overload set. *) tvSetValue (typesTypeVar n, tvValue tyVar); madeList := Matched {old = tyVar, new = n} :: !madeList; n end ) end fun copyTypeVar (atyp as TypeVar tyVar) = if tvLevel tyVar <> generalisable then atyp (* Not generalisable. *) else (* Unbound, overload set or flexible record *) let val newTv = tvs atyp in (* If we have a type variable pointing to a flexible record we have to copy the type pointed at by the variable. *) case tvValue tyVar of valu as LabelledType _ => tvSetValue (typesTypeVar newTv, copyType (valu, copyTypeVar, fn t => t)) | _ => (); newTv end | copyTypeVar atyp = atyp val copied = (* Only process type variables. Return type constructors unchanged. *) copyType (atyp, copyTypeVar, fn t => t (*copyTCons*)) in (copied, ! madeList) end (* generaliseTypes *); (* Exported wrapper for generaliseTypes. *) fun generalise atyp = let val (t, newMatch) = generaliseTypes (atyp, fn _ => NONE) fun makeResult(Matched{new, old}) = {value=new, equality=tvEquality old, printity=tvPrintity old} in (t, List.map makeResult newMatch) end; (* Return the original polymorphic type variables. *) fun getPolyTypeVars(atyp, map) = let val (_, newMatch) = generaliseTypes (atyp, map) in List.map (fn(Matched{old, ...}) => old) newMatch end; fun generaliseWithMap(atyp, map) = let val (t, newMatch) = generaliseTypes (atyp, map) fun makeResult(Matched{new, old}) = {value=new, equality=tvEquality old, printity=tvPrintity old} in (t, List.map makeResult newMatch) end (* Find the argument type which gives this result when the constructor is applied. If we have, for example, a value of type int list and we have discovered that this is a "::" node we have to work back by comparing the type of "::" ('a * 'a list -> 'a list) to find the argument of the constructor (int * int list) and hence how to print it. (Actually "list" is treated specially). *) fun constructorResult (FunctionType{arg, result=TypeConstruction{args, ...}}, typeArgs) = let val matches = ListPair.zip(List.map typesTypeVar args, typeArgs) fun getArg tv = case List.find(fn (atv, _) => sameTv(tv, atv)) matches of SOME (_, ty) => SOME ty | NONE => NONE in #1 (generaliseTypes(arg, getArg)) end | constructorResult _ = raise InternalError "Not a function type" (* If we have a type construction which is an alias for another type we construct the alias by first instantiating all the type variables and then copying the type. *) fun makeEquivalent (atyp, args) = case tcIdentifier atyp of TypeId{idKind=TypeFn(typeArgs, typeResult), ...} => let val matches = ListPair.zip(typeArgs, args) fun getArg tv = case List.find(fn (atv, _) => sameTv(tv, atv)) matches of SOME (_, ty) => SOME ty | NONE => NONE in #1 (generaliseTypes(typeResult, getArg)) end | TypeId _ => raise InternalError "makeEquivalent: Not a type function" (* Look for the occurrence of locally declared datatypes in the type of a value. *) fun checkForEscapingDatatypes(ty: types, errorFn: string->unit) : unit = let fun checkTypes (typ: types) (ok: bool) : bool = case typ of TypeConstruction {constr, args, ...} => if tcIsAbbreviation constr then (* May be an alias for a type that contains a local datatype. *) foldType checkTypes (makeEquivalent (constr, args)) ok else if ok then ( case tcIdentifier constr of TypeId{access=Local{addr, ...}, ...} => if !addr < 0 then ( errorFn("Type of expression contains local datatype (" ^ tcName constr ^") outside its definition."); false ) else true | _ => true (* Could we have a "selected" entry with a local datatype? *) ) else false | _ => ok in foldType checkTypes ty true; () end (* This 3-valued logic is used because in a few cases we may not be sure if equality testing is allowed. If we have 2 mutually recursive datatypes t = x of s | ... and s = z of t we would first examine "t", find that it uses "s", look at "s", find that refers back to "t". To avoid infinite recursion we return the result that equality "maybe" allowed for "t" and hence for "s". However we may find that the other constructors for "t" do not allow equality and so equality will not be allowed for "s" either. *) datatype tri = Yes (* 3-valued logic *) | No | Maybe; (* Returns a flag saying if equality testing is allowed for values of the given type. "equality" is used both to generate the code for a specific call of equality e.g. (a, b, c) = f(x), and to generate the equality operation for a type when it is declared. In the latter case type variables may be parameters which will be filled in later e.g. type 'a list = nil | op :: of ('a * 'a list). "search" is a function which looks up constructors in mutually recursive type declarations. "lookupTypeVar" deals with type variables. If they represent parameters to a type declaration equality checking will be allowed. If we are unifying this type to an equality type variable they will be unified to new equality type variables. Otherwise equality is not allowed. *) fun equality (ty, search, lookupTypeVar) : tri = let (* Can't use foldT because it is not monotonic (equality on ref 'a is allowed). *) (* Returns Yes only if equality testing is allowed for all types in the list. *) fun eqForList ([], soFar) = soFar | eqForList (x::xs, soFar) = case equality (x, search, lookupTypeVar) of No => No | Maybe => eqForList (xs, Maybe) | Yes => eqForList (xs, soFar); in case eventual ty of TypeVar tyVar => (* The type variable may point to a flexible record or an overload set or it may be the end of the chain. If this is a labelled record we have to make sure that any fields we add also admit equality. lookupTypeVar makes the type variable an equality type so that any new fields are checked for equality but we also have to call "equality" to check the existing fields. *) if tvEquality tyVar then Yes else ( case tvValue tyVar of lab as LabelledType _ => ( case lookupTypeVar tyVar of No => No | _ => equality (lab, search, lookupTypeVar) ) | _ => lookupTypeVar tyVar ) | FunctionType {...} => No (* No equality on function types! *) | TypeConstruction {constr, args, ...} => if isUndefined constr then No else if tcIsAbbreviation constr then (* May be an alias for a type that allows equality. *) equality (makeEquivalent (constr, args), search, lookupTypeVar) (* ref - Equality is permitted on refs of all types *) (* The Definition of Standard ML says that ref is the ONLY type constructor which is treated in this way. The standard basis library says that other mutable types such as array should also work this way. *) else if isPointerEqType(tcIdentifier constr) then Yes (* Others apart from ref and real *) else if tcEquality constr (* Equality allowed. *) then eqForList (args, Yes) (* Must be allowed for all the args *) else let (* Not an alias. - Look it up. *) val s = search (tcIdentifier constr); in if s = No then No else eqForList (args, s) end (* TypeConstruction *) | LabelledType {recList, ...} => (* Record equality if all subtypes are (ignore frozen!) *) (* TODO: Avoid copying the list? *) eqForList (map (fn{typeof, ...}=>typeof) recList, Yes) | OverloadSet _ => (* This should not happen because all overload sets should be pointed to by type variables and so should be handled in the TypeVar case. *) raise InternalError "equality - Overloadset found" | BadType => No | EmptyType => No (* shouldn't occur *) end (* When a datatype is declared we test to see if equality is allowed. The types are mutually recursive so value constructors of one type may take arguments involving values of any of the others. *) fun computeDatatypeEqualities(types: typeConstrSet list, boundIdEq) = let datatype state = Processed of tri (* Already processed or processing. *) | NotSeen of typeConstrSet list (* Value is list of constrs. *); (* This table tells us, for each type constructor, whether it definitely admits equality, definitely does not or whether we have yet to look at it. *) fun isProcessed (Processed _) = true | isProcessed _ = false; fun stateProcessed (Processed x) = x | stateProcessed _ = raise Match; fun stateNotSeen (NotSeen x) = x | stateNotSeen _ = raise Match; val {enter:typeId * state -> unit,lookup} = mapTable sameTypeId; (* Look at each of the constructors in the list. Equality testing is only allowed if it is allowed for each of the alternatives. *) fun constrEq _ [] soFar = soFar (* end of list - all o.k. *) | constrEq constructor (h :: t) soFar = (* The constructor may be a constant e.g. datatype 'a list = nil | ... or a function e.g. datatype 'a list = ... cons of 'a * 'a list. *) if not (isFunctionType (valTypeOf h)) (* Constant *) then constrEq constructor t soFar (* Go on to the next. *) else let (* Function - look at the argument type. *) (* Equality is allowed for any type-variable. The only type variables allowed are parameters to the datatype so if we have a type variable then equality is allowed for this datatype. *) val eq = equality (#arg (typesFunctionType (valTypeOf h)), genEquality, fn _ => Yes); in if eq = No then (* Not allowed. *) No else (* O.k. - go on to the next. *) constrEq constructor t (if eq = Maybe then Maybe else soFar) end (* constrEq *) (* This procedure checks to see if equality is allowed for this datatype. *) and genEquality constructorId = let (* Look it up to see if we have already done it. It may fail because we may have constructors that do not admit equality. *) val thisState = case (lookup constructorId, constructorId) of (SOME inList, _) => inList | (NONE, TypeId{idKind = Bound{offset, ...}, ...}) => Processed(if boundIdEq offset then Yes else No) | _ => Processed No in if isProcessed thisState then stateProcessed thisState (* Have either done it already or are currently doing it. *) else (* notSeen - look at it now. *) let (* Equality is allowed for this datatype only if all of them admit it. There are various other alternatives but this is what the standard says. If the "name" is rigid (free) we must not grant equality if it is not already there although that is not an error. *) (* Set the state to "Maybe". This prevents infinite recursion. *) val () = enter (constructorId, Processed Maybe); val eq = List.foldl (fn (cons, t) => if t = No then No else constrEq cons (tsConstructors cons) t) Yes (stateNotSeen thisState); in (* Set the state we have found if it is "yes" or "no". If it is maybe we have a recursive reference which appears to admit equality, but may not. E.g. if we have datatype t = A of s | B of int->int and s = C of t if we start processing "t" we will go on to "s" and do that before returning to "t". It is only later we find that "t" does not admit equality. If we get "Maybe" as the final result when all the recursion has been unwound we can set the result to "yes", but any intermediate "Maybe"s have to be done again. *) enter (constructorId, if eq = Maybe then thisState else Processed eq); eq end end (* genEquality *); in (* If we have an eqtype we set it to true, otherwise we set all of them to "notSeen" with the constructor as value. *) List.app (fn dec as TypeConstrSet(decCons, _) => let (* If we have two datatypes which share we may already have one in the table. We have to link them together. *) val tclist = case lookup (tcIdentifier decCons) of NONE => [dec] | SOME l => let val others = stateNotSeen l val newList = dec :: others; in (* If any of these are already equality types (i.e. share with an eqtype) then they all must be. *) if tcEquality decCons orelse tcEquality (tsConstr(hd others)) then List.app (fn d => tcSetEquality (tsConstr d, true)) newList else (); newList end in enter (tcIdentifier decCons, NotSeen tclist) end) types; (* Apply genEquality to each element of the list. *) List.app (fn TypeConstrSet(constructor, _) => let val constructorId = tcIdentifier constructor; val eqForCons = genEquality constructorId; in (* If the result is "Maybe" it involves a recursive reference, but the rest of the type allows equality. The type admits equality. *) if eqForCons = No then () (* Equality not allowed *) else ( (* Turn on equality. *) enter (constructorId, Processed Yes); tcSetEquality (constructor, true) ) end) types end (* computeDatatypeEqualities *); datatype matchResult = SimpleError of types * types * string | TypeConstructorError of types * types * typeConstrs * typeConstrs (* Type matching algorithm for both unification and signature matching. *) (* The mapping has now been moved out of here. Instead when signature matching the target signature is copied before this is called which means that this process is now symmetric. There may be some redundant tests left in here. *) fun unifyTypes(Atype : types, Btype : types) : matchResult option = let (* Get the result in here. This isn't very ML-like but it greatly simplifies converting the code. *) val matchResult: matchResult option ref = ref NONE fun matchError error = (* Only report one error. *) case matchResult of ref (SOME _) => () | r => r := SOME error fun cantMatch(alpha, beta, text) = matchError(SimpleError(alpha, beta, text)) fun match (Atype : types, Btype : types) : unit = let (* Check two records/tuples and return the combined type. *) fun unifyRecords (rA as {recList=typAlist, fullList = gA}, rB as {recList=typBlist, fullList = gB}, typA : types, typB : types) : types = let val typAFrozen = recordIsFrozen rA and typBFrozen = recordIsFrozen rB fun matchLabelled ([], []) = [] (* Something left in bList - this is fine if typeA is not frozen. e.g. (a: s, b: t) will match (a: s, ...) but not just (a:s). *) | matchLabelled ([], bList as {name=bName, ...} :: _) = ( if typAFrozen then cantMatch (typA, typB, "(Field " ^ bName ^ " missing)") else (); bList (* return the remainder of the list *) ) | matchLabelled (aList as {name=aName, ...} :: _, []) = (* Something left in bList *) ( if typBFrozen then cantMatch (typA, typB, "(Field " ^ aName ^ " missing)") else (); aList (* the rest of aList *) ) | matchLabelled (aList as ((aVal as {name=aName,typeof=aType})::aRest), bList as ((bVal as {name=bName,typeof=bType})::bRest)) = (* both not nil - look at the names. *) let val order = compareLabels (aName, bName); in if order = 0 (* equal *) then (* same name - must be unifiable types *) ( (* The result is (either) one of these with the rest of the list. *) match (aType, bType); aVal :: matchLabelled (aRest, bRest) ) else if order < 0 (* aName < bName *) then (* The entries in each list are in order so this means that this entry is not in bList. If the typeB is frozen this is an error. *) if typBFrozen (* Continue with the entry removed. *) then (cantMatch (typA, typB, "(Field " ^ aName ^ " missing)"); aList) else aVal :: matchLabelled (aRest, bList) else (* aName > bName *) if typAFrozen then (cantMatch (typA, typB, "(Field " ^ bName ^ " missing)"); bList) else bVal :: matchLabelled (aList, bRest) end (* not nil *); (* Return the combined list. Only actually used if both are flexible. *) val result = if typAFrozen andalso typBFrozen andalso List.length typAlist <> List.length typBlist then (* Don't attempt to unify the fields if we have the wrong number of items. If we've added or removed an item from a tuple e.g. a function with multiple arguments, it's more useful to know this than to get unification errors on fields that don't match. *) (cantMatch (typA, typB, "(Different number of fields)"); []) else matchLabelled (typAlist, typBlist) fun lastFlex(FlexibleList(ref(r as FlexibleList _))) = lastFlex r | lastFlex(FlexibleList r) = SOME r | lastFlex(FieldList _) = NONE in if typAFrozen then (if typBFrozen then () else valOf(lastFlex gB) := gA; typA) else if typBFrozen then (valOf(lastFlex gA) := gB; typB) else let (* We may have these linked already in which case we shouldn't do anything. *) val lastA = valOf(lastFlex gA) and lastB = valOf(lastFlex gB) in if lastA = lastB then () else let val genericFields = FieldList(map #name result, false) in (* If these are both flexible we have link all the generics together so that if we freeze any one of them they all get frozen. *) lastA := genericFields; lastB := FlexibleList lastA end; LabelledType {recList = result, fullList = gA} end end (* unifyRecords *); (* Sets a type variable to a value. - Checks that the type variable we are assigning does not occur in the expression we are about to assign to it. Such cases can occur if we have infinitely-typed expressions such as fun a. a::a where a has type 'a list list ... Also propagates the level information of the type variable. Now also deals with flexible records. *) fun assign (var, t) = let (* Mapped over the type to be assigned. *) (* Returns "false" if it is safe to make the assignment. Sorts out imperative type variables and propagates level information. N.B. It does not propagate equality status. The reason is that if we are unifying ''a with 'b ref, the 'b does NOT become an equality type var. In all other cases it would. *) fun occursCheckFails _ true = true | occursCheckFails ty false = let val t = eventual ty in case t of TypeVar tvar => let (* The level is the minimum of the two, and if we are unifying with an equality type variable we must make this into one. *) val minLev = Int.min (tvLevel var, tvLevel tvar) val oldValue = tvValue tvar in if tvLevel tvar <> minLev then (* If it is nonunifiable we cannot make its level larger. *) if tvNonUnifiable tvar then cantMatch (Atype, Btype, "(Type variable is free in surrounding scope)") else let (* Must make a new type var with the right properties *) (* This type variable may be a flexible record, in which case we have to save the record and put it on the new type variable. We have to do this for the record itself so that new fields inherit the correct status and also for any existing fields. *) val newTv = mkTypeVar (minLev, tvEquality tvar, false, tvPrintity tvar) in tvSetValue (typesTypeVar newTv, oldValue); tvSetValue (tvar, newTv) end else (); (* Safe if vars are different but we also have to check any flexible records. *) occursCheckFails oldValue (sameTv (tvar, var)) end | TypeConstruction {args, constr, ...} => (* If this is a type abbreviation we have to expand this before processing any arguments. We mustn't process arguments that are not actually used. *) if tcIsAbbreviation constr then occursCheckFails(makeEquivalent (constr, args)) false else List.foldr (fn (t, v) => occursCheckFails t v) false args | FunctionType {arg, result} => occursCheckFails arg false orelse occursCheckFails result false | LabelledType {recList,...} => List.foldr (fn ({ typeof, ... }, v) => occursCheckFails typeof v) false recList | _ => false end val varVal = tvValue var (* Current value of the variable to be set. *) local (* We need to process any type abbreviations before applying the occurs check. The type we're assigning could boil down to the same type variable we're trying to assign. This doesn't breach the occurs check. *) fun followVarsAndTypeFunctions t = case eventual t of ev as TypeConstruction{constr, args, ...} => if tcIsAbbreviation constr then followVarsAndTypeFunctions(makeEquivalent (constr, args)) else ev | ev => ev in val finalType = followVarsAndTypeFunctions t end (* We may actually have the same type variable after any type abbreviations have been followed. *) val reallyTheSame = case finalType of TypeVar tv => sameTv (tv, var) | _ => false in (* start of "assign" *) case varVal of LabelledType _ => (* Flexible record. Check that the records are compatible. *) match (varVal, t) | OverloadSet _ => (* OverloadSet. Check that the sets match. This is only in the case where t is something other than an overload set since we remove the overload set from a variable when unifying two sets. *) match (varVal, t) | _ => (); if reallyTheSame then () (* Don't apply the occurs check or check for non-unifiable. *) (* If this type variable was put in explicitly then it can't be assigned to something else. (We have already checked for the type variables being the same). *) else if tvNonUnifiable var then cantMatch (Atype, Btype, "(Cannot unify with explicit type variable)") else if occursCheckFails finalType false then cantMatch (Atype, Btype, "(Type variable to be unified occurs in type)") else let (* Occurs check succeeded. *) fun canMkEqTv (tvar : typeVarForm) : tri = (* Turn it into an equality type var. *) if tvEquality tvar then Yes (* If it is nonunifiable we cannot make it into an equality type var. *) else if tvNonUnifiable tvar then No else (* Must make a new type var with the right properties *) let (* This type variable may be a flexible record or an overload set, in which case we have to save the record and put it on the new type variable. We have to do both because we have to ensure that the existing fields in the flexible record admit equality and ALSO that any additional fields we may add by unification with other records also admit equality. *) val newTv = mkTypeVar (tvLevel tvar, true, false, tvPrintity tvar) val oldValue = tvValue tvar in tvSetValue (tvar, newTv); (* If this is an overloaded type we must remove any types that don't admit equality. *) case oldValue of OverloadSet{typeset} => let (* Remove any types which do not admit equality. *) fun filter [] = [] | filter (h::t) = if tcEquality h then h :: filter t else filter t in case filter typeset of [] => No | [constr] => ( (* Turn a singleton into a type construction. *) tvSetValue (typesTypeVar newTv, mkTypeConstruction(tcName constr, constr, nil, [])); Yes ) | newset => ( tvSetValue (typesTypeVar newTv, OverloadSet{typeset=newset}); Yes ) end | _ => (* Labelled record or unbound variable. *) ( tvSetValue (typesTypeVar newTv, oldValue); Yes ) end in (* If we are unifying a type with an equality type variable we must ensure that equality is allowed for that type. This will turn most type variables into equality type vars. *) if tvEquality var andalso equality (t, fn _ => No, canMkEqTv) = No then cantMatch (Atype, Btype, "(Requires equality type)") (* TODO: This can result in an unhelpful message if var is bound to a flexible record since there is no indication in the printed type that the flexible record is an equality type. It would be improved if we set the value to be EmptyType. At least then the type variable would be printed which would be an equality type. --- Adding the "Requires equality type" should improve things. *) else (); (* Propagate the "printity" status. This is probably not complete but doesn't matter too much since this is a Poly extension. *) if tvPrintity var then let fun makePrintity(TypeVar tv) _ = ( if tvPrintity tv then () else case tvValue tv of (* If it's an overload set we don't need to do anything. This will eventually be a monotype. *) OverloadSet _ => () | oldValue => let (* Labelled record or unbound variable. *) val newTv = mkTypeVar (tvLevel tv, tvEquality tv, tvNonUnifiable tv, true) in tvSetValue(tv, newTv); (* Put this on the chain if it's a labelled record. *) tvSetValue (typesTypeVar newTv, oldValue) end ) | makePrintity _ _ = () in foldType makePrintity t () end else (); (* Actually make the assignment. It doesn't matter if var is a labelled record, because t will be either a fixed record or a combination of the fields of var and t. Likewise if var was previously an overload set this may replace the set by a single type construction. *) (* If we have had an error don't make the assignment. At the very least it could prevent us producing useful error information and it could also result in unnecessary consequential errors. *) case !matchResult of NONE => tvSetValue (var, t) | SOME _ => () end end (* assign *); (* First find see if typeA and typeB are unified to anything already, and get the end of a list of "flexibles". *) val tA = eventual Atype and tB = eventual Btype in (* start of "match" *) if isUndefinedType tA orelse isUndefinedType tB then () (* If either of these was an undefined type constructor don't try to match. TODO: There are further tests below for this which are now redundant. *) else case (tA, tB) of (BadType, _) => () (* If either is an error don't try to match *) | (_, BadType) => () | (TypeVar typeAVar, TypeVar typeBVar) => (* Unbound type variable, flexible record or overload set. *) let (* Even if this is a one-way match we can allow type variables in the typeA to be instantiated to anything in the typeB. *) val typeAVal = tvValue typeAVar; (* We have two unbound type variables or flex. records. *) in if sameTv (typeAVar, typeBVar) (* same type variable? *) then () else (* no - assign one to the other *) if tvNonUnifiable typeAVar (* If we have a nonunifiable type variable we want to assign the typeB to it. If the typeB is nonunifiable as well we will get an error message. *) then assign (typeBVar, tA) else let (* If they are both flexible records we first set the typeB to the union of the records, and then set the typeA to that. In that way we propagate properties such as equality and level between the two variables. *) val typBVal = tvValue typeBVar in case (typeAVal, typBVal) of (LabelledType recA, LabelledType recB) => ( (* Turn these back into simple type variables to save checking the combined record against the originals when we make the assignment. (Would be safe but redundant). *) tvSetValue (typeBVar, emptyType); tvSetValue (typeAVar, emptyType); assign (typeBVar, unifyRecords (recA, recB, typeAVal, typBVal)); assign (typeAVar, tB) ) | (OverloadSet{typeset=setA}, OverloadSet{typeset=setB}) => let (* The lists aren't ordered so we just have to go through by hand. *) fun intersect(_, []) = [] | intersect(a, H::T) = if isInSet(H, a) then H::intersect(a, T) else intersect(a, T) val newSet = intersect(setA, setB) in case newSet of [] => cantMatch (Atype, Btype, "(Incompatible overloadings)") | _ => ( tvSetValue (typeBVar, emptyType); tvSetValue (typeAVar, emptyType); (* I've changed this from OverloadSet{typeset=newset} to use mkOverloadSet. The main reason was that it fixed a bug which resulted from a violation of the assumption that "equality" would not be passed an overload set except when pointed to by a type variable. It also removed the need for a separate test for singleton sets since mkOverloadSet deals with them. DCJM 1/9/00. *) assign (typeBVar, mkOverloadSet newSet); assign (typeAVar, tB) ) end | (EmptyType, _) => (* A is not a record or an overload set. *) assign (typeAVar, tB) | (_, EmptyType) => (* A is a record but B isn't *) assign (typeBVar, tA) (* typeB is ordinary type var. *) | _ => (* Bad combination of labelled record and overload set *) cantMatch (Atype, Btype, "(Incompatible types)") end end | (TypeVar typeAVar, _) => (* typeB is not a type variable so set typeA to typeB.*) (* Be careful if this is a non-unifiable type variable being matched to the special case of the identity type-construction. *) ( if tvNonUnifiable typeAVar orelse (case tvValue typeAVar of OverloadSet _ => true | _ => false) then ( case tB of TypeConstruction {constr, args, ...} => if isUndefined constr orelse not (tcIsAbbreviation constr) then ( case tB of TypeConstruction {constr, args, ...} => if isUndefined constr orelse not (tcIsAbbreviation constr) then assign (typeAVar, tB) else match(tA, eventual (makeEquivalent (constr, args))) | _ => assign (typeAVar, tB) ) else match(tA, eventual (makeEquivalent (constr, args))) | _ => assign (typeAVar, tB) ) else assign (typeAVar, tB) ) | (_, TypeVar typeBVar) => (* and typeA is not *) ( (* We have to check for the special case of the identity type-construction. *) if tvNonUnifiable typeBVar orelse (case tvValue typeBVar of OverloadSet _ => true | _ => false) then ( case tA of TypeConstruction {constr, args, ...} => if isUndefined constr orelse not (tcIsAbbreviation constr) then ( case tB of TypeVar tv => (* This will fail if we are matching a signature because the typeB will be non-unifiable. *) assign (tv, tA) (* set typeB to typeA *) | typB => match (tA, typB) ) else match(eventual (makeEquivalent (constr, args)), tB) | _ => ( case tB of TypeVar tv => (* This will fail if we are matching a signature because the typeB will be non-unifiable. *) assign (tv, tA) (* set typeB to typeA *) | typB => match (tA, typB) ) ) else ( case tB of TypeVar tv => (* This will fail if we are matching a signature because the typeB will be non-unifiable. *) assign (tv, tA) (* set typeB to typeA *) | typB => match (tA, typB) ) ) | (TypeConstruction({constr = tACons, args=tAargs, ...}), TypeConstruction ({constr = tBCons, args=tBargs, ...})) => ( (* We may have a number of possibilities here. a) If tA is an alias we simply expand it out and recurse (even if tB is the same alias). e.g. if we have string t where type 'a t = int*'a we expand string t into int*string and try to unify that. b) map it and see if the result is an alias. -- NOW REMOVED c) If tB is a type construction and it is an alias we expand that e.g. unifying "int list" and "int t" where type 'a t = 'a list (particularly common in signature/structure matching.) d) Finally we try to unify the stamps and the arguments. *) if isUndefined tACons orelse isUndefined tBCons then () (* If we've had an undefined type constructor don't try to check further. *) else if tcIsAbbreviation tACons (* Candidate is an alias - expand it. *) then match (makeEquivalent (tACons, tAargs), tB) else if tcIsAbbreviation tBCons then match (tA, makeEquivalent (tBCons, tBargs)) else if tcIsAbbreviation tBCons (* If the typeB is an alias it must be expanded. *) then match (tA, makeEquivalent (tBCons, tBargs)) else if sameTypeId (tcIdentifier tACons, tcIdentifier tBCons) then let (* Same type constructor - do the arguments match? *) fun matchLists [] [] = () | matchLists (a::al) (b::bl) = ( match (a, b); matchLists al bl ) | matchLists _ _ = (* This should only happen as a result of a different error. *) cantMatch (Atype, Btype, "(Different numbers of arguments)") in matchLists tAargs tBargs end (* When we have different type constructors, especially two with the same name, we try to produce more information. *) else matchError(TypeConstructorError(tA, tB, tACons, tBCons)) ) | (OverloadSet {typeset}, TypeConstruction {constr=tBCons, args=tBargs, ...}) => (* The candidate is an overloaded type and the target is a type construction. *) ( if not (isUndefined tBCons orelse not (tcIsAbbreviation tBCons)) then match (tA, makeEquivalent (tBCons, tBargs)) else if isUndefined tBCons then () else if tcIsAbbreviation tBCons then match (tA, makeEquivalent (tBCons, tBargs)) else (* See if the target type is among those in the overload set. *) if null tBargs (* Must be a nullary type constructor. *) andalso isInSet(tBCons, typeset) then () (* ok. *) (* Overload sets arise primarily with literals such as "1" and it's most likely that the error is a mismatch between int and another type rather than that the user assumed that the literal was overloaded on a type it actually wasn't. *) else case preferredOverload typeset of NONE => cantMatch (tA, tB, "(Different type constructors)") | SOME prefType => matchError( TypeConstructorError( mkTypeConstruction (tcName prefType, prefType,[], []), tB, prefType, tBCons)) ) | (TypeConstruction {constr=tACons, args=tAargs, ...}, OverloadSet {typeset}) => ( if not (isUndefined tACons orelse not (tcIsAbbreviation tACons)) then match (makeEquivalent (tACons, tAargs), tB) (* We should never find an overload set as the target for a signature match but it is perfectly possible for tB to be an overload set when unifying two types. *) else if null tAargs andalso isInSet(tACons, typeset) then () (* ok. *) else case preferredOverload typeset of NONE => cantMatch (tA, tB, "(Different type constructors)") | SOME prefType => matchError( TypeConstructorError( tA, mkTypeConstruction (tcName prefType, prefType,[], []), tACons, prefType)) ) | (OverloadSet _ , OverloadSet _) => raise InternalError "Unification: OverloadSet/OverloadSet" (* (OverloadSet , OverloadSet) should not occur because that should be handled in the (TypeVar, TypeVar) case. *) | (TypeConstruction({constr = tACons, args=tAargs, ...}), _) => if not (isUndefined tACons orelse not (tcIsAbbreviation tACons)) (* Candidate is an alias - expand it. *) then match (makeEquivalent (tACons, tAargs), tB) else (* typB not a construction (but typeA is) *) cantMatch (tA, tB, "(Incompatible types)") | (_, TypeConstruction {constr=tBCons, args=tBargs, ...}) => (* and typeA is not. *) (* May have a type equivalence e.g. "string t" matches int*string if type 'a t = int * 'a . Alternatively we may be matching a structure to a signature where the signature says "type t" and the structure contains "type t = int->int" (say). We need to set the type in the signature to int->int. *) if not (isUndefined tBCons orelse not (tcIsAbbreviation tBCons)) then match (tA, makeEquivalent (tBCons, tBargs)) else if isUndefined tBCons then () else if tcIsAbbreviation tBCons then match (tA, makeEquivalent (tBCons, tBargs)) else cantMatch (tB, tA, "(Incompatible types)") | (FunctionType {arg=typAarg, result=typAres, ...}, FunctionType {arg=typBarg, result=typBres, ...}) => ( (* must be unifiable functions *) (* In principle it doesn't matter whether we unify arguments or results first but it could affect the error messages. Is this the best way to do it? *) match (typAarg, typBarg); match (typAres, typBres) ) | (EmptyType, EmptyType) => () (* This occurs only with exceptions - empty means no argument *) | (LabelledType recA, LabelledType recB) => (* Unify the records, but discard the result because at least one of the records is frozen. *) (unifyRecords (recA, recB, tA, tB); ()) | _ => cantMatch (tA, tB, "(Incompatible types)") end (* match *) in match (Atype, Btype); ! matchResult end (* unifyTypes *) (* Turn a result from matchTypes into a pretty structure so that it can be included in a message. *) fun unifyTypesErrorReport (_, alphaTypeEnv, betaTypeEnv, what) = let fun reportError(SimpleError(alpha: types, beta: types, reason)) = (* This previously used a single type variable sequence for both types. It may be that this is needed to make sensible error messages. *) PrettyBlock(3, false, [], [ PrettyString ("Can't " ^ what (* "match" if a signature, "unify" if core lang. *)), PrettyBreak (1, 0), display (alpha, 1000 (* As deep as necessary *), alphaTypeEnv), PrettyBreak (1, 0), PrettyString "to", PrettyBreak (1, 0), display (beta, 1000 (* As deep as necessary *), betaTypeEnv), PrettyBreak (1, 0), PrettyString reason ]) | reportError(TypeConstructorError(alpha: types, beta: types, alphaCons, betaCons)) = let fun expandedTypeConstr(ty, tyEnv, tyCons) = let fun lastPart name = #second(splitString name) (* Print the type which includes the type constructor name with as much additional information as we can. *) fun printWithDesc{ location, name, description } = PrettyBlock(3, false, [], [ display (ty, 1000, tyEnv) ] @ (if lastPart name = lastPart(tcName tyCons) then [] else [ PrettyBreak(1, 0), PrettyString "=", PrettyBreak(1, 0), PrettyBlock(0, false, [ContextLocation location], [PrettyString name]) ] ) @ (if description = "" then [] else [ PrettyBreak(1, 0), PrettyBlock(0, false, [ContextLocation location], [PrettyString ("(*" ^ description ^ "*)")]) ] ) ) in case tcIdentifier tyCons of TypeId { description, ...} => printWithDesc description end in PrettyBlock(3, false, [], [ PrettyString ("Can't " ^ what (* "match" if a signature, "unify" if core lang. *)), PrettyBreak (1, 0), expandedTypeConstr(alpha, alphaTypeEnv, alphaCons), PrettyBreak (1, 0), PrettyString (if what = "unify" then "with" else "to"), PrettyBreak (1, 0), expandedTypeConstr(beta, betaTypeEnv, betaCons), PrettyBreak (1, 0), PrettyString "(Different type constructors)" ]) end in reportError end (* Given a function type returns the first argument if the function takes a tuple otherwise returns the only argument. Extended to include the case where the argument is not a function in order to work properly for overloaded literals. *) fun firstArg(FunctionType{arg= LabelledType { recList = {typeof, ...} ::_, ...}, ...}) = eventual typeof | firstArg(FunctionType{arg, ...}) = eventual arg | firstArg t = t (* Returns an instance of an overloaded function using the supplied list of type constructors for the overloading. *) fun generaliseOverload(t, constrs, isConverter) = let (* Returns the result type of a function. *) fun getResult(FunctionType{result, ...}) = eventual result | getResult _ = raise InternalError "getResult - not a function"; val arg = if isConverter then getResult t else firstArg t in case arg of TypeVar tv => let (* The argument should be a type variable, possibly set to an empty overload set. This should be replaced by the current overload set in the copied function type. *) val newSet = mkOverloadSet constrs val (t, _) = generaliseTypes(t, fn old => if sameTv(old, tv) then SOME newSet else NONE) in (t, [newSet]) end | _ => raise InternalError "generaliseOverload - arg is not a type var" end (* Prints out a type constructor e.g. type 'a fred = 'a * 'a or datatype 'a joe = bill of 'a list | mary of 'a * int or simply type 'a abs if the type is abstract. *) fun displayTypeConstrsWithMap ( TypeConstrSet( TypeConstrs{identifier=TypeId{idKind=TypeFn(args, result), ...}, name, ...}, []), depth, typeEnv, sigMap) = (* Type function *) if depth <= 0 then PrettyString "..." else let val typeV = varNameSequence () (* Local sequence for this binding. *) in PrettyBlock (3, false, [], PrettyString "type" :: PrettyBreak (1, 0) :: printTypeVars (args, depth, typeV) @ [ PrettyString (#second(splitString name)), PrettyBreak(1, 0), PrettyString "=", PrettyBreak(1, 0), tDisp(result, depth-1, typeV, typeEnv, sigMap) ] ) end | displayTypeConstrsWithMap (TypeConstrSet(tCons, [] (* No constructors *)), depth, _, _) = (* Abstract type or type in a signature. *) if depth <= 0 then PrettyString "..." else PrettyBlock (3, false, [], PrettyString ( if tcEquality tCons then "eqtype" else "type") :: PrettyBreak (1, 0) :: printTypeVars (tcTypeVars tCons, depth, varNameSequence ()) @ [PrettyString (#second(splitString(tcName tCons)))] ) | displayTypeConstrsWithMap (TypeConstrSet(tCons as TypeConstrs{name, locations, ...}, tcConstructors), depth, typeEnv, sigMap) = (* It has constructors - datatype declaration *) if depth <= 0 then PrettyString "..." else let val typeV = varNameSequence () (* Construct a ('a, 'b, 'c) tyCons construction for the result types of each of the constructors. N.B. We use the original type constructors because they have the appropriate equality type properties. datatype 'a t = A of 'a is not the same as ''a t = A of ''a. *) val typeVars = tcTypeVars tCons val typeResult = mkTypeConstruction(name, tCons, map TypeVar typeVars, locations) (* Print a single constructor (blocked) *) fun pValConstr (first, name, typeOf, depth) = let val (t, _) = generalise typeOf val firstBreak = PrettyBreak (1, if first then 2 else 0) in case t of FunctionType { arg, result} => let (* Constructor with an argument. The constructor "type" is the argument. We have to unify the result type of the function with the ('a, 'b, 'c) tyCons type so that we get the correct type variables in the argument. We just print the argument of the function. *) val _ = unifyTypes(result, typeResult) in [ firstBreak, PrettyBlock (0, false, [], PrettyBlock (0, false, [], (if first then PrettyBreak (0, 2) else PrettyBlock (0, false, [], [PrettyString "|", PrettyBreak(1, 2)]) ) :: (if depth <= 0 then [PrettyString "..."] else [ PrettyString name, PrettyBreak (1, 4), PrettyString "of"]) ) :: (if depth > 0 then [ PrettyBreak (1, 4), (* print the type as a single block of output *) tDisp (arg, depth - 1, typeV, typeEnv, sigMap) ] else []) ) ] end | _ => [ firstBreak, PrettyBlock (0, false, [], [if first then PrettyBreak (0, 2) else PrettyBlock (0, false, [], [PrettyString "|", PrettyBreak(1, 2)]), PrettyString (if depth <= 0 then "..." else name)] ) ] end (* Print a sequence of constructors (unblocked) *) fun pValConstrRest ([], _ ): pretty list = [] | pValConstrRest (H :: T, depth): pretty list = if depth < 0 then [] else pValConstr (false, valName H, valTypeOf H, depth) @ pValConstrRest (T, depth - 1) fun pValConstrList ([], _ ) = PrettyString "" (* shouldn't occur *) | pValConstrList (H :: T, depth) = PrettyBlock (2, true, [], pValConstr (true, valName H, valTypeOf H, depth) @ pValConstrRest (T, depth - 1) ) in PrettyBlock(0, false, [], [ PrettyBlock(0, false, [], PrettyString "datatype" :: PrettyBreak (1, 2) :: printTypeVars (typeVars, depth, typeV) @ [ PrettyString(#second(splitString(tcName tCons))), PrettyBreak(1, 0), PrettyString "=" ] ), pValConstrList (tcConstructors, depth - 1) ] ) end (* displayTypeConstrsWithMap *) fun displayTypeConstrs (tCons : typeConstrSet, depth : FixedInt.int, typeEnv) : pretty = displayTypeConstrsWithMap(tCons, depth, typeEnv, NONE) (* Return a type constructor from an overload. If there are several (i.e. the overloading has not resolved to a single type) it returns the "best". This is called in the third pass so it should never be called if there is not at least one type that is possible. *) fun typeConstrFromOverload(f, _) = let fun prefType(TypeVar tvar) = ( (* If we still have an overload set that's because it has not reduced to a single type. In ML 97 we default to int, real, word, char or string in that order. This works correctly for overloading literals so long as the literal conversion functions are correctly installed. *) case tvValue tvar of OverloadSet{typeset} => let (* If we accept this type we have to freeze the overloading to this type. I'm not happy about doing this here but it seems the easiest solution. *) fun freezeType tcons = ( tvSetValue(tvar, mkTypeConstruction(tcName tcons, tcons, [], [])); tcons ) in case preferredOverload typeset of SOME tycons => freezeType tycons | NONE => raise InternalError "typeConstrFromOverload: No matching type" end | _ => raise InternalError "typeConstrFromOverload: No matching type" (* Unbound or flexible record. *) ) | prefType(TypeConstruction{constr, args, ...}) = if not (tcIsAbbreviation constr) then constr (* Generally args will be nil in this case but in the special case of looking for an equality function for 'a ref or 'a array it may not be. *) else prefType (makeEquivalent (constr, args)) | prefType _ = raise InternalError "typeConstrFromOverload: No matching type" in prefType(firstArg(eventual f)) end; (* Return the result type of a function. Also used to test if the value is a function type. *) fun getFnArgType t = case eventual t of FunctionType {arg, ... } => SOME arg | _ => NONE (* Assigns type variables to variables with generalisation permitted if their level is at least that of the current level. In ML90 mode this produced an error message for any top-level free imperative type variables. We don't do that in ML97 because it is possible that another declaration may "freeze" the type variable before the composite expression reaches the top level. *) fun allowGeneralisation (t, level, nonExpansive, lex, location, moreInfo, typeEnv) = let fun giveError(s1: string, s2: string) = let (* Use a single sequence. *) val vars : typeVarForm -> string = varNameSequence (); open DEBUG val parameters = debugParams lex val errorDepth = getParameter errorDepthTag parameters in reportError lex { hard = true, location = location, message = PrettyBlock (3, false, [], [ PrettyString s1, PrettyBreak (1, 0), tDisp (t, errorDepth, vars, typeEnv, NONE), PrettyBreak (1, 0), PrettyString s2 ] ), context = SOME(moreInfo ()) } end local open DEBUG val parameters = debugParams lex in val checkOverloadFlex = getParameter narrowOverloadFlexRecordTag parameters end fun general t (genArgs as (showError, nonExpansive)) = case eventual t of TypeVar tvar => let val argSet = if tvLevel tvar >= level andalso tvLevel tvar <> generalisable andalso (case tvValue tvar of OverloadSet _ => false | _ => true) then let (* Make a new generisable type variable, except that type variables in an expansive context cannot be generalised. We also don't generalise if this is an overload set. The reason for that is that it allows us to get overloading information from the surrounding context. e.g. let fun f x y = x+y in f 2.0 end. An alternative would be take the default type (in this case int). DCJM 1/9/00. *) val nonCopiable = not nonExpansive val newLevel = if nonCopiable then level-1 else generalisable (* copiable *); val isOk = (* If the type variable has top-level scope then we have a free type variable. We only want to generate this message once even if we have multiple type variables.*) (* If the type variable is non-unifiable and the expression is expansive then we have an error since this will have to be a monotype. *) if tvNonUnifiable tvar andalso nonCopiable andalso showError then ( giveError("Type", "includes a free type variable"); false ) else showError; (* It may be a flexible record so we have to transfer the record to the new variable. *) val newTypeVar = makeTv {value=tvValue tvar, level=newLevel, equality=tvEquality tvar, nonunifiable=if nonCopiable then (tvNonUnifiable tvar) else false, printable=tvPrintity tvar} in tvSetValue (tvar, TypeVar newTypeVar); (* If we are using the "narrow" context for overloading and flexible records we should apply this here. Otherwise it is dealt with in the next pass when we have the full program context. *) case (checkOverloadFlex, tvValue tvar) of (true, LabelledType _) => giveError("Type", "is an unresolved flexible record") | (true, OverloadSet {typeset, ...}) => ( (* Set this to the "preferred" type. Typically this is "int" but for overloaded literals (e.g. 0w0) it could be something else. *) case preferredOverload typeset of SOME tycons => tvSetValue(tvar, mkTypeConstruction(tcName tycons, tycons, [], [])) | NONE => raise InternalError "general: No matching type" ) | _ => (); (isOk, nonExpansive) end else genArgs in general (tvValue tvar) argSet (* Process any flexible record. *) end | TypeConstruction {args, constr, ...} => (* There is a pathological case here. If we have a type equivalence which contains type variables that do not occur on the RHS (e.g. type 'a t = int) then we generalise over them even with an expansive expression. This is because the semantics treats type abbreviations as type functions and so any type variables that are eliminated by the function application do not appear in the "type" that the semantics applies to the expression. *) if tcIsAbbreviation constr then let val (r1, _) = general(makeEquivalent (constr, args)) genArgs (* Process any arguments that have not been processed in the equivalent. *) val (r2, _) = List.foldr (fn (t, v) => general t v) (r1, true) args in (r2, nonExpansive) end else List.foldr (fn (t, v) => general t v) genArgs args | FunctionType {arg, result} => general arg (general result genArgs) | LabelledType {recList,...} => List.foldr (fn ({ typeof, ... }, v) => general typeof v) genArgs recList | _ => genArgs in general t (true, nonExpansive); () end (* end allowGeneralisation *); (* Check for free type variables at the top level. Added for ML97. This replaces the test in allowGeneralisation above and is applied to all top-level values including those in structures and functors. *) (* I've changed this from giving an error message, which prevented the code from evaluating, to giving a warning and setting the type variables to unique type variables. That allows, for example, fun f x = raise x; f Subscript; to work. DCJM 8/3/01. *) fun checkForFreeTypeVariables(valName: string, ty: types, lex: lexan, printAndEqCode) : unit = let (* Generate new names for the type constructors. *) val count = ref 0 fun genName num = (if num >= 26 then genName (num div 26 - 1) else "") ^ String.str (Char.chr (num mod 26 + Char.ord #"a")); fun checkTypes (TypeVar tvar) () = if isEmpty(tvValue tvar) andalso tvLevel tvar = 1 then (* The type variable is unbound (specifically, not an overload set) and it is not generic i.e. it must have come from an expansive expression. *) let val name = "_" ^ genName(!count) val _ = count := !count + 1; val declLoc = location lex (* Not correct but OK for the moment. *) val declDescription = { location = declLoc, name = name, description = "Constructed from a free type variable." } val tCons = makeTypeConstructor (name, [], makeFreeId(0, Global(printAndEqCode()), tvEquality tvar, declDescription), [DeclaredAt declLoc]); val newVal = mkTypeConstruction(name, tCons, [], []) in warningMessage(lex, location lex, concat["The type of (", valName, ") contains a free type variable. Setting it to a unique monotype."]); tvSetValue (tvar, newVal) end else () | checkTypes _ () = () in foldType checkTypes ty (); () end (* Returns true if a type constructor permits equality. *) fun permitsEquality constr = if tcIsAbbreviation constr then typePermitsEquality( mkTypeConstruction (tcName constr, constr, List.map TypeVar (tcTypeVars constr), [])) else tcEquality constr and typePermitsEquality ty = equality (ty, fn _ => No, fn _ => Yes) <> No (* See if a type abbreviation or "where type" has the form type t = s or type 'a t = 'a s etc and so is simply giving a new name to the type constructor. If it is it then checks that the type constructor used (s in this example) is just a simple type name. *) fun typeNameRebinding(typeArgs, typeResult): typeId option = let fun eqTypeVar(TypeVar ta, tb) = sameTv (ta, tb) | eqTypeVar _ = false in case typeResult of TypeConstruction {constr, args, ... } => if not (ListPair.allEq eqTypeVar(args, typeArgs)) then NONE else ( case tcIdentifier constr of TypeId{idKind=TypeFn _, ...} => NONE | tId => SOME tId ) | _ => NONE end (* Returns the number of the entry in the list. Used to find out the location of fields in a labelled record for expressions and pattern matching. Assumes that the label appears in the list somewhere. *) fun entryNumber (label, LabelledType{recList, ...}) = let (* Count up the list. *) fun entry ({name, ...}::l) n = if name = label then n else entry l (n + 1) | entry [] _ = raise Match in entry recList 0 end | entryNumber (label, TypeVar tvar) = entryNumber (label, tvValue tvar) | entryNumber (label, TypeConstruction{constr, ...}) = (* Type alias *) entryNumber (label, tcEquivalent constr) | entryNumber _ = raise InternalError "entryNumber - not a record" (* Size of a labelled record. *) fun recordWidth (LabelledType{recList, ...}) = length recList | recordWidth (TypeVar tvar) = recordWidth (tvValue tvar) | recordWidth (TypeConstruction{constr, ...}) = (* Type alias *) recordWidth (tcEquivalent constr) | recordWidth _ = raise InternalError "entryNumber - not a record" fun recordFieldMap f (LabelledType{recList, ...}) = List.map (f o (fn {typeof, ...} => typeof)) recList | recordFieldMap f (TypeVar tvar) = recordFieldMap f (tvValue tvar) | recordFieldMap f (TypeConstruction{constr, ...}) = recordFieldMap f (tcEquivalent constr) | recordFieldMap _ _ = raise InternalError "entryNumber - not a record" (* Unify two type variables which would otherwise be non-unifiable. Used when we have found a local type variable with the same name as a global one. *) fun linkTypeVars (a, b) = let val ta = typesTypeVar (eventual(TypeVar a)); (* Must both be type vars. *) val tb = typesTypeVar (eventual(TypeVar b)); in (* Set the one with the higher level to point to the one with the lower, so that the effective level is the lower. *) if (tvLevel ta) > (tvLevel tb) then tvSetValue (ta, TypeVar b) else tvSetValue (tb, TypeVar a) end; (* Set its level by setting it to a new type variable. *) fun setTvarLevel (typ, level) = let val tv = typesTypeVar (eventual(TypeVar typ)); (* Must be type var. *) in tvSetValue (tv, mkTypeVar (level, tvEquality tv, true, tvPrintity tv)) end; (* Construct the least general type from a list of types. This is used after type checking to try to remove polymorphism from local values. It takes the list of actual uses of the value, usually a function, and removes any unnecessary polymorphism. This is particularly the case if the function involves a flexible record, where the unspecified fields are treated as polymorphic, but where the function is actually applied to a records which are monomorphic. *) fun leastGeneral [] = EmptyType (* Never used? *) (* Don't use this at the moment - see the comment on TypeVar below. Also the comment on TypeConstruction for local datatypes. *) (* | leastGeneral [oneType] = oneType *)(* Just one - this is it. *) | leastGeneral(firstType::otherTypes): types = let fun canonical (typ as TypeVar tyVar) = ( case tvValue tyVar of EmptyType => typ | OverloadSet _ => let val constr = typeConstrFromOverload(typ, false) in mkTypeConstruction(tcName constr, constr, [], []) end | t => canonical t ) | canonical (typ as TypeConstruction { constr, args, ...}) = if tcIsAbbreviation constr (* Handle type abbreviations directly *) then canonical(makeEquivalent (constr, args)) else typ | canonical typ = typ (* Take the head of the each argument list and extract the least general. Then process the tail. It's an error if each element of the list does not contain the same number of items. *) fun leastArgs ([]::_) = [] | leastArgs (args as _::_) = leastGeneral(List.map hd args) :: leastArgs (List.map tl args) | leastArgs _ = raise Empty in case canonical firstType of (*typ as *)TypeVar _(*tv*) => let (*fun sameTypeVar(TypeVar tv1) = sameTv(tv, tv1) | sameTypeVar _ = false*) in (* If they are all the same type variable return that otherwise return a new generalisable type variable. They may all be equal if we always apply this function to a value whose type is a polymorphic type in the function that contains all these uses. *) (* Temporarily, at least, create a new type var in this case. If we have a polymorphic function that is only used inside another polymorphic function but isn't declared inside it, if we use the caller's type variable here the call won't be recognised as polymorphic. *) (*if List.all sameTypeVar otherTypes then typ else*) mkTypeVar(generalisable, false, false, false) end | TypeConstruction{ constr, args, name, locations, ...} => ( (* There is a potential problem if the datatype is local including if it was constructed in a functor. Almost always it will have been declared after the polymorphic function but if it happens not to have been we could set a polymorphic function to a type that doesn't exist yet. To avoid this we don't allow a local datatype here and instead fall back to the polymorphic case. *) case tcIdentifier constr of thisConstrId as TypeId{access=Global _, ...} => let val argLength = List.length args (* This matches if it is an application of the same type constructor. *) fun getTypeConstrs(TypeConstruction{constr, args, ...}) = if sameTypeId(thisConstrId, tcIdentifier constr) andalso List.length args = argLength then SOME args else NONE | getTypeConstrs _ = NONE val allArgs = List.mapPartial (getTypeConstrs o canonical) otherTypes in if List.length allArgs = List.length otherTypes then TypeConstruction{constr=constr, name=name, locations=locations, args = leastArgs(args :: allArgs)} else (* At least one of these wasn't the same type constructor. *) mkTypeVar(generalisable, false, false, false) end | _ => mkTypeVar(generalisable, false, false, false) ) | FunctionType{ arg, result } => let fun getFuns(FunctionType{arg, result}) = SOME(arg, result) | getFuns _ = NONE val argResults = List.mapPartial (getFuns o canonical) otherTypes in if List.length argResults = List.length otherTypes then let val (args, results) = ListPair.unzip argResults in FunctionType{arg=leastGeneral(arg::args), result = leastGeneral(result::results)} end else (* At least one of these wasn't a function. *) mkTypeVar(generalisable, false, false, false) end | LabelledType (r as {recList=firstRec, fullList}) => if recordIsFrozen r then let (* This matches if all the field names are the same. Extract the types. *) fun nameMatch({name=name1: string, ...}, {name=name2, ...}) = name1 = name2 fun getRecords(LabelledType{recList, ...}) = if ListPair.allEq nameMatch (firstRec, recList) then SOME(List.map #typeof recList) else NONE | getRecords _ = NONE val argResults = List.mapPartial (getRecords o canonical) otherTypes in if List.length argResults = List.length otherTypes then let (* Use the names from the first record (they all are the same) to build a new record. *) val argTypes = leastArgs(List.map #typeof firstRec :: argResults) fun recreateRecord({name, ...}, types) = {name=name, typeof=types} val newList = ListPair.map recreateRecord(firstRec, argTypes) in LabelledType{recList=newList, fullList=fullList } end else (* At least one of these wasn't a record. *) mkTypeVar(generalisable, false, false, false) end else (* At this stage the record should be frozen if the program is correct but if it isn't we could have a flexible record which we report elsewhere. *) mkTypeVar(generalisable, false, false, false) | _ => (* May arise if there's been an error. *) mkTypeVar(generalisable, false, false, false) end (* Test if this is floating point i.e. the "real" type. We could include abbreviations of real as well but it's probably not worth it. *) datatype floatKind = FloatDouble | FloatSingle local val realId = tcIdentifier realConstr and floatId = tcIdentifier floatConstr fun isFloatId constr = let val id = tcIdentifier constr in if sameTypeId(id, realId) then SOME FloatDouble else if sameTypeId(id, floatId) then SOME FloatSingle else NONE end in fun isFloatingPt(TypeConstruction{args=[], constr, ...}) = isFloatId constr | isFloatingPt(OverloadSet {typeset, ...}) = ( case preferredOverload typeset of SOME t => isFloatId t (* real only. float is never preferred. *) | NONE => NONE ) | isFloatingPt(TypeVar tv) = isFloatingPt (tvValue tv) | isFloatingPt _ = NONE end fun checkDiscard(t: types, lex: lexan): string option = let open DEBUG val checkLevel = getParameter reportDiscardedValuesTag (debugParams lex) fun isUnit(LabelledType{recList=[], ...}) = true (* Unit is actually an empty record *) | isUnit(TypeConstruction{ constr as TypeConstrs{identifier=TypeId{idKind=TypeFn _, ...}, ...}, args, ...}) = isUnit(makeEquivalent(constr, args)) | isUnit(TypeVar _) = true (* Allow unbound type vars *) | isUnit _ = false fun isAFunction(FunctionType _) = true | isAFunction(TypeConstruction{ constr as TypeConstrs{identifier=TypeId{idKind=TypeFn _, ...}, ...}, args, ...}) = isAFunction(makeEquivalent(constr, args)) | isAFunction _ = false in case checkLevel of 1 => if isAFunction (eventual t) then SOME "A function value is being discarded." else NONE | 2 => if isUnit (eventual t) then NONE else SOME "A non unit value is being discarded." | _ => NONE end structure Sharing = struct type types = types and values = values and typeId = typeId and structVals = structVals and typeConstrs= typeConstrs and typeConstrSet=typeConstrSet and typeParsetree = typeParsetree and locationProp = locationProp and pretty = pretty and lexan = lexan and ptProperties = ptProperties and typeVarForm = typeVarForm and codetree = codetree and matchResult = matchResult and generalMatch = generalMatch end end (* TYPETREE *); diff --git a/mlsource/MLCompiler/VALUE_OPS.ML b/mlsource/MLCompiler/VALUE_OPS.ML index 6eb1ecc2..de28adb1 100644 --- a/mlsource/MLCompiler/VALUE_OPS.ML +++ b/mlsource/MLCompiler/VALUE_OPS.ML @@ -1,1321 +1,1321 @@ (* Copyright (c) 2000 Cambridge University Technical Services Limited - Modified David C.J. Matthews 2008-9, 2013, 2015-16. + Modified David C.J. Matthews 2008-9, 2013, 2015-16, 2020. This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License version 2.1 as published by the Free Software Foundation. This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU Lesser General Public License for more details. You should have received a copy of the GNU Lesser General Public License along with this library; if not, write to the Free Software Foundation, Inc., 51 Franklin St, Fifth Floor, Boston, MA 02110-1301 USA *) (* Title: Operations on global and local values. Author: Dave Matthews, Cambridge University Computer Laboratory Copyright Cambridge University 1986 *) functor VALUE_OPS ( structure LEX : LEXSIG; structure CODETREE : CODETREESIG structure STRUCTVALS : STRUCTVALSIG; structure TYPESTRUCT : TYPETREESIG structure PRINTTABLE : PRINTTABLESIG structure UNIVERSALTABLE: sig type universal = Universal.universal type univTable val app: (string * universal -> unit) -> univTable -> unit end; structure DEBUG : DEBUGSIG structure MISC : sig exception InternalError of string; (* compiler error *) exception Conversion of string (* string to int conversion failure *) val quickSort : ('a -> 'a -> bool) -> 'a list -> 'a list end; structure PRETTY : PRETTYSIG structure ADDRESS : AddressSig structure UTILITIES : sig val splitString: string -> { first:string,second:string } end; structure COPIER: COPIERSIG structure TYPEIDCODE: TYPEIDCODESIG structure DATATYPEREP: DATATYPEREPSIG sharing STRUCTVALS.Sharing = TYPESTRUCT.Sharing = LEX.Sharing = PRETTY.Sharing = COPIER.Sharing = CODETREE.Sharing = PRINTTABLE = ADDRESS = UNIVERSALTABLE = MISC = TYPEIDCODE.Sharing = DATATYPEREP.Sharing ) : VALUEOPSSIG = (*****************************************************************************) (* VALUEOPS functor body *) (*****************************************************************************) struct open MISC; open PRETTY; open LEX; open CODETREE; open TYPESTRUCT; (* Open this first because unitType is in STRUCTVALS as well. *) open Universal; (* for tag etc. *) open STRUCTVALS; open PRINTTABLE; open DEBUG; open ADDRESS; open UTILITIES; open TYPEIDCODE open COPIER open DATATYPEREP (* Functions to construct the values. *) fun mkGconstr (name, typeof, code, nullary, constrs, location) = makeValueConstr (name, typeof, nullary, constrs, Global code, location); (* Global variable *) fun mkGvar (name, typeOf, code, locations) : values = Value{ name = name, typeOf = typeOf, access = Global code, class = ValBound, locations = locations, references = NONE, instanceTypes=NONE }; (* Local variable - Generated by the second pass. *) local fun makeLocalV class (name, typeOf, locations) = Value{ name = name, typeOf = typeOf, access = Local {addr = ref ~1 (* Must be set later *), level = ref baseLevel}, class = class, locations = locations, references = makeRef(), instanceTypes=SOME(ref []) }; in val mkValVar = makeLocalV ValBound and mkPattVar = makeLocalV PattBound end (* Value in a local structure or a functor argument. May be simple value, exception or constructor. *) fun mkSelectedVar (Value { access = Formal addr, name, typeOf, class, locations, ...}, Struct{access=sAccess, ...}, openLocs) = (* If the argument is "formal" set the base to the base structure. *) let (* If the base structure is a constant do the selection now. This is redundant unless we're being called from PolyML.NameSpace.Structures.contents. *) val access = case sAccess of Global code => Global(mkInd (addr, code)) | _ => Selected{addr=addr, base=sAccess} in Value{name=name, typeOf=typeOf, class=class, access=access, locations=openLocs @ locations, references = NONE, instanceTypes=NONE} end | mkSelectedVar (Value { access = Global code, name, typeOf, class, locations, ...}, _, openLocs) = (* Global: We need to add the location information. *) Value{name=name, typeOf=typeOf, class=class, access=Global code, locations=openLocs @ locations, references = NONE, instanceTypes=NONE} | mkSelectedVar(selected, _, _) = selected (* Overloaded? *); (* Construct a global exception. *) fun mkGex (name, typeof, code, locations) = Value{ name = name, typeOf = typeof, access = Global code, class = Exception, locations = locations, references = NONE, instanceTypes=NONE } (* Construct a local exception. *) fun mkEx (name, typeof, locations) = Value{ name = name, typeOf = typeof, access = Local{addr = ref 0, level = ref baseLevel}, class = Exception, locations=locations, references = NONE, instanceTypes=NONE } (* Locations in exception packets. In order to have a defined ordering of the fields, when we put the location in an exception packet we use this datatype rather than the "location" type. *) (* *) datatype RuntimeLocation = NoLocation | SomeLocation of (* file: *) string * (*startLine:*) int * (*startPosition:*) int * (*endLine:*) int * (*endPosition:*) int fun codeLocation({file="", startLine=0, startPosition=0, ...}) = mkConst(toMachineWord NoLocation) (* No useful information *) | codeLocation({file, startLine, startPosition, endLine, endPosition}) = mkConst(toMachineWord(file, startLine, startPosition, endLine, endPosition)) (*****************************************************************************) (* Look-up functions. *) (* These are used locally and also exported to INITIALISE to be used in PolyML.NameSpace.Structures.contents. *) fun makeSelectedValue( Value{ name, typeOf, access, class, locations, ... }, baseStruct as Struct{signat=Signatures { typeIdMap, ...}, name=baseName, ...}) = let fun copyId(TypeId{idKind=Bound{ offset, ...}, ...}) = SOME(typeIdMap offset) | copyId _ = NONE val copiedType = copyType (typeOf, fn x => x, fn tcon => copyTypeConstr (tcon, copyId, fn x => x, fn s => baseName^"."^s)) val baseLoc = case List.find (fn DeclaredAt _ => true | _ => false) locations of SOME (DeclaredAt loc) => [StructureAt loc] | _ => [] in mkSelectedVar ( Value{ name=name, typeOf=copiedType, access=access, class=class, locations=locations, references = NONE, instanceTypes=NONE }, baseStruct, baseLoc) end fun makeSelectedStructure( Struct {signat, access, name=structName, locations, ...}, Struct {signat=Signatures { typeIdMap, firstBoundIndex, ...}, access=baseAccess, ...}) = let val Signatures { name=sigName, tab, typeIdMap = childMap, locations=sigLocs, ... } = signat (* We need to apply the map from the parent structure to the child. *) val copiedSig = makeSignature(sigName, tab, firstBoundIndex, sigLocs, composeMaps(childMap, typeIdMap), []) (* Convert Formal access to Selected and leave the others (Global?). If this is Formal but the base structure is global do the selection now. This is only needed if we're called from PolyML.NameSpace.Structures.contents. *) val newAccess = case (access, baseAccess) of (Formal sel, Global code) => Global(mkInd(sel, code)) | (Formal sel, baseAccess) => Selected { addr = sel, base = baseAccess } | (access, _) => access (* If we have a DeclaredAt location for the structure use this as the StructureAt.*) val baseLoc = case List.find (fn DeclaredAt _ => true | _ => false) locations of SOME (DeclaredAt loc) => [StructureAt loc] | _ => [] in Struct { name = structName, signat = copiedSig, access = newAccess, locations = baseLoc @ locations} end fun makeSelectedType(typeConstr, Struct { signat=Signatures { typeIdMap, ...}, name, ...}) = fullCopyDatatype(typeConstr, typeIdMap, name^".") (* Look up a structure. *) fun lookupStructure (kind, {lookupStruct:string -> structVals option}, name, errorMessage) = let val {first = prefix, second = suffix} = splitString name; val strLookedUp = if prefix = "" then lookupStruct suffix else case lookupStructure ("Structure", {lookupStruct=lookupStruct}, prefix, errorMessage) of NONE => NONE (* Already reported *) | SOME(baseStruct as Struct { signat=Signatures { tab, ... }, ...}) => let (* Look up the first part in the structure environment. *) val Env{lookupStruct, ...} = makeEnv tab in case lookupStruct suffix of SOME foundStruct => SOME(makeSelectedStructure(foundStruct, baseStruct)) | NONE => NONE end in case strLookedUp of SOME s => SOME s | NONE => (* Not declared? *) (errorMessage (kind ^ " (" ^ suffix ^ ") has not been declared" ^ (if prefix = "" then "" else " in structure " ^ prefix)); NONE) end fun mkEnv x = let val Env e = makeEnv x in e end (* Look up a structure but ignore the access. This is used in sharing constraints where we're only interested in the signature. *) (* It's simpler to use the common code for this. *) fun lookupStructureAsSignature (lookupStruct, name, errorMessage) = lookupStructure("Structure", { lookupStruct = lookupStruct}, name, errorMessage) (* Look up a value, possibly in a structure. If it is in a structure we may have to apply a selection. *) fun lookupValue (kind, {lookupVal,lookupStruct}, name, errorMessage) = let val {first = prefix, second = suffix} = splitString name; val found = if prefix = "" then lookupVal suffix (* Look up the first part in the structure environment. *) else case lookupStructure ("Structure", {lookupStruct=lookupStruct}, prefix, errorMessage) of NONE => SOME undefinedValue | SOME (baseStruct as Struct { signat=Signatures { tab, ...}, ...}) => ( case #lookupVal (mkEnv tab) suffix of SOME foundValue => SOME(makeSelectedValue(foundValue, baseStruct)) | NONE => NONE ) in case found of SOME v => v | NONE => (* Not declared? *) ( errorMessage (kind ^ " (" ^ suffix ^ ") has not been declared" ^ (if prefix = "" then "" else " in structure " ^ prefix)); undefinedValue ) end fun lookupTyp ({lookupType,lookupStruct}, name, errorMessage) = let val {first = prefix, second = suffix} = splitString name; val found = if prefix = "" then lookupType suffix else (* Look up the first part in the structure environment. *) case lookupStructure ("Structure", {lookupStruct=lookupStruct}, prefix, errorMessage) of NONE => SOME(TypeConstrSet(undefConstr, [])) | SOME (baseStruct as Struct { signat=Signatures { tab, ...}, ...}) => ( case #lookupType (mkEnv tab) suffix of SOME typeConstr => SOME(makeSelectedType(typeConstr, baseStruct)) | NONE => NONE ) in case found of SOME v => v | NONE => (* Not declared? *) ( errorMessage ("Type constructor" ^ " (" ^ suffix ^ ") has not been declared" ^ (if prefix = "" then "" else " in structure " ^ prefix)); TypeConstrSet(undefConstr, []) ) end (* Printing. *) (* Print a value given its type. *) fun printValueForType (value:machineWord, types, depth): pretty = let (* Constuct printer code applied to the argument and the depth. Code-generate and evaluate it. *) (* If this is polymorphic apply it to a dummy set of instance types. This may happen if we have val it = NONE at the top level. The equality attributes of the type variables must match so that this works correctly with justForEqualityTypes set. *) val addrs = ref 0 (* Make local declarations for any type values. *) local fun mkAddr n = !addrs before (addrs := !addrs + n) in val typeVarMap = TypeVarMap.defaultTypeVarMap(mkAddr, baseLevel) end val dummyTypes = List.map(fn tv => {value=TYPESTRUCT.unitType, equality=tvEquality tv, printity=false}) (getPolyTypeVars(types, fn _ => NONE)) val polyCode = applyToInstance(dummyTypes, baseLevel, typeVarMap, fn _ => mkConst value) val printerCode = mkEval( printerForType(types, baseLevel, typeVarMap), [mkTuple[polyCode, mkConst(toMachineWord depth)]]) val pretty = RunCall.unsafeCast( valOf(evalue(genCode(CODETREE.mkEnv(TypeVarMap.getCachedTypeValues typeVarMap, printerCode), [], !addrs)()))) in pretty end (* These are used to display the declarations made. *) fun displayFixStatus(FixStatus(name, f)): pretty = let open PRETTY val status = case f of Nonfix => PrettyString "nonfix" | Infix prec => PrettyBlock(0, false, [], [ PrettyString "infix", PrettyBreak (1, 0), PrettyString (Int.toString prec) ]) | InfixR prec => PrettyBlock(0, false, [], [ PrettyString "infixr", PrettyBreak (1, 0), PrettyString (Int.toString prec) ]) in PrettyBlock (0, false, [], [status, PrettyBreak (1, 0), PrettyString name]) end (* Returns the declaration location as the location for the context. *) fun getLocation locations = case List.find(fn DeclaredAt _ => true | _ => false) locations of SOME(DeclaredAt loc) => [ContextLocation loc] | _ => [] (* Displays value as a block, with no external formatting. This is used at the top level but it can be applied to values extracted with #lookup globalNameSpace. That can include constructors and overloaded functions. *) fun displayValues (Value{name, typeOf, class, access, locations, ...}, depth: FixedInt.int, nameSpace, sigMap): pretty = let (* Create the "val X =" part. *) fun valPart (valOrCons, isColon) = let (* If we're putting in a colon we don't need a space after an alphanumeric id but we do if it's symbolic. *) val isAlphaNumeric = let val first = String.sub(name, 0) in Char.isAlpha first orelse first = #"'" end val space = if isColon andalso isAlphaNumeric then 0 else 1 val equOrColon = if isColon then ":" else "=" in PrettyBlock (0, false, [], [ PrettyString valOrCons, PrettyBreak (1, 0), PrettyBlock(0, false, getLocation locations, [PrettyString name]), PrettyBreak (space, 0), PrettyString equOrColon ] ) end val typeEnv = (* Environment to check for type constructors. *) { lookupType = #lookupType nameSpace, lookupStruct = #lookupStruct nameSpace} in if depth <= 0 then PrettyString "..." else case class of ValBound => let (* In nearly all cases if we have Global code we will have a constant. There was one case where "!" was actually a Lambda that hadn't been code-generated. *) val value = case access of Global code => evalue code | _ => NONE val start = case value of SOME v => [ valPart("val", false), PrettyBreak (1, 0), printValueForType (v, typeOf, depth), PrettyString ":" ] | _ => [ valPart("val", true) ] in PrettyBlock (3, false, [], start @ [ PrettyBreak (1, 0), displayWithMap (typeOf, depth, typeEnv, sigMap) ]) end | Exception => (* exceptions *) PrettyBlock (0, false, [], PrettyBlock (0, false, [], [ PrettyString "exception", PrettyBreak (1, 0), PrettyBlock(0, false, getLocation locations, [PrettyString name]) ] ) :: ( case getFnArgType typeOf of NONE => [] | SOME excType => [ PrettyBreak (1, 1), PrettyString "of", PrettyBreak (1, 3), displayWithMap (excType, depth, typeEnv, sigMap) ] ) ) | Constructor _ => (* This can only occur with #lookupVal *) PrettyBlock (3, false, [], [ valPart("constructor", true), PrettyBreak (1, 0), displayWithMap (typeOf, depth, typeEnv, sigMap) ]) | PattBound => (* Can this ever occur? *) PrettyBlock (3, false, [], [ valPart("val", true), PrettyBreak (1, 0), displayWithMap (typeOf, depth, typeEnv, sigMap) ]) end (* Print global values. This is passed through the bootstrap and used in the debugger. *) fun printValues (Value{typeOf, class, access, ...}, depth) = case (class, access) of (ValBound, Global code) => printValueForType (valOf(evalue code), typeOf, depth) | _ => PrettyString "" (* Probably shouldn't occur. *) (* Prints "sig ... end" as a block, with no external formatting *) fun displaySig (Signatures{tab, typeIdMap, ...}, depth : FixedInt.int, _ : int, { lookupType, lookupStruct, ...}, sigMap: (int-> typeId) option) : pretty = let (* Construct an environment for the types. *) val Env { lookupType = strType, lookupStruct = strStr, ...} = makeEnv tab (* Construct a map for types. *) val innerMap = case sigMap of NONE => SOME typeIdMap | SOME outerMap => SOME(composeMaps(typeIdMap, outerMap)) val compositeEnv = { lookupType = fn s => case strType s of NONE => lookupType s | SOME t => SOME (t, innerMap), lookupStruct = fn s => case strStr s of NONE => lookupStruct s | SOME s => SOME (s, innerMap) } val typeEnv: printTypeEnv = { lookupType = #lookupType compositeEnv, lookupStruct = #lookupStruct compositeEnv } fun displaySpec (_, value) : pretty list = if (tagIs signatureVar value) then (* Not legal ML97 *) [ PrettyBreak(1,2), displaySignatures (tagProject signatureVar value, depth - 1, compositeEnv)] else if (tagIs structVar value) then [ PrettyBreak(1,2), displayStructures (tagProject structVar value, depth - 1, compositeEnv, innerMap)] else if (tagIs typeConstrVar value) then [ PrettyBreak(1,2), displayTypeConstrsWithMap (tagProject typeConstrVar value, depth, typeEnv, innerMap) ] else if (tagIs valueVar value) then let (* Only print variables. Constructors are printed with their type. *) val value = tagProject valueVar value; in case value of Value{class = Constructor _, ...} => [] | _ => [ PrettyBreak(1,2), (* We lookup the infix status and any exception in the global environment only. Infix status isn't a property of a structure and it's too much trouble to look up exceptions in the structure. *) displayValues (value, depth, compositeEnv, innerMap) ] end else if (tagIs fixVar value) then (* Not legal ML97 *) [ PrettyBreak(1,2), displayFixStatus (tagProject fixVar value) ] else [] (* end displaySpec *) in PrettyBlock (0, true, [], PrettyString "sig" :: ( ( if depth <= 1 (* If the depth is 1 each of the calls to displaySpec will print "..." so we replace them all by a single "..." here. *) then [PrettyBreak (1, 0), PrettyString "..."] else let val declist = ref nil : (string * universal) list ref fun addToList nv = declist := nv :: !declist (* For the moment order them by name. We may change this to order primarily by kind and secondarily by name. *) fun order (s1: string, _) (s2: string, _) = s1 > s2 in (* Put all the entries into a list. *) UNIVERSALTABLE.app addToList tab; (* Sort the list and print it. *) List.foldl (fn (a, l) => displaySpec a @ l) [] (quickSort order (!declist)) end ) @ [PrettyBreak (1, 0), PrettyString "end"] ) ) end (* displaySig *) (* Print: signature S = sig .... end *) and displaySignatures (str as Signatures{locations, name, ...}, depth : FixedInt.int, nameSpace) : pretty = if depth <= 0 then PrettyString "..." else PrettyBlock(0, false, [], [ PrettyBlock(0, false, [], [ PrettyString "signature", PrettyBreak(1, 0), PrettyBlock(0, false, getLocation locations, [PrettyString name]), PrettyBreak(1, 0), PrettyString "=" ] ), PrettyBreak (1, 2), displaySig (str, depth, 1, nameSpace, NONE) ]) (* print structure in a block (no external spacing) *) and displayStructures (Struct{name, locations, signat, ...}, depth, nameSpace, sigMap): pretty = if depth <= 0 then PrettyString "..." else PrettyBlock (0, false, [], [ PrettyBlock(0, false, [], [ PrettyString "structure", PrettyBreak(1, 0), PrettyBlock(0, false, getLocation locations, [PrettyString name]), PrettyBreak(0, 0), PrettyString ":" ] ), PrettyBreak(1, 2), displayNamedSig(signat, depth - 1, 1, nameSpace, sigMap) ]) (* Internal function for printing structures and functors. If a signature has a name print the name rather than the contents. *) and displayNamedSig(sign as Signatures{name = "", ...}, depth, space, nameSpace, sigMap) = displaySig (sign, depth, space, nameSpace, sigMap) | displayNamedSig(Signatures{name, ...}, _, _, _, _) = PrettyString name fun displayFunctors (Functor{ name, locations, arg, result, ...}, depth, nameSpace) = if depth <= 0 then PrettyString "..." else let val arg as Struct { name = argName, signat as Signatures { tab = argTab, ... }, ...} = arg val argEntries = (if argName <> "" then [ PrettyBlock(0, false, [], [PrettyString argName, PrettyBreak(0, 0), PrettyString ":"]), PrettyBreak(1, 2) ] else []) @ [ displayNamedSig (signat, depth - 1, 0, nameSpace, NONE), PrettyBreak(0, 0), PrettyString "):", PrettyBreak(1, 0) ] (* Include the argument structure name in the type environment. *) val argEnv = if argName = "" then let val Env { lookupType=lt, lookupStruct=ls, ...} = makeEnv argTab in { lookupType = fn s => case lt s of NONE => #lookupType nameSpace s | SOME t => SOME(t, NONE), lookupStruct = fn s => case ls s of NONE => #lookupStruct nameSpace s | SOME s => SOME(s, NONE) } end else { lookupType = #lookupType nameSpace, lookupStruct = fn s => if s = argName then SOME(arg, NONE) else #lookupStruct nameSpace s } in PrettyBlock (0, false, [], [ PrettyBlock(0, false, [], [ PrettyBlock(0, false, [], [ PrettyString "functor", PrettyBreak(1, 0), PrettyBlock(0, false, getLocation locations, [PrettyString name]), PrettyBreak(1, 0), PrettyString "(" ]), PrettyBreak(0, 2), PrettyBlock(0, false, [], argEntries) ]), PrettyBreak(0, 2), displayNamedSig (result, depth - 1, 1, argEnv, NONE) ] ) end (* Exported version. *) val displayValues = fn (value, depth, nameSpace) => displayValues (value, depth, nameSpace, NONE) and displayStructures = fn (str, depth, nameSpace) => displayStructures (str, depth, nameSpace, NONE) (* Code-generation. *) (* Code-generate the values. *) fun codeStruct (Struct{access, ...}, level) = (* Global structures have no code value. Instead the values are held in the values of the signature. *) codeAccess (access, level) and codeAccess (Global code, _) = code | codeAccess (Local{addr=ref locAddr, level=ref locLevel}, level) = mkLoad (locAddr, level, locLevel) (* Argument or local *) | codeAccess (Selected{addr, base}, level) = (* Select from a structure. *) mkInd (addr, codeAccess (base, level)) | codeAccess _ = raise InternalError "No access" (*****************************************************************************) (* datatype access functions *) (*****************************************************************************) (* Get the appropriate instance of an overloaded function. If the overloading has not resolved to a single type it finds the preferred type if possible (i.e. int for most overloadings, but possibly real, word, string or char for conversion functions.) *) fun getOverloadInstance(name, instance, isConv): codetree*string = let val constr = typeConstrFromOverload(instance, isConv) in (getOverload(name, constr, fn _ => raise InternalError "getOverloadInstance: Missing"), tcName constr) end (* This is only used in addPrettyPrint. There's no point in producing a lot of detailed information. *) fun checkPPType (instanceType, matchType, fnName, lex, location, moreInfo) = case unifyTypes (instanceType, matchType) of NONE => () | SOME error => let open DEBUG val parameters = LEX.debugParams lex val errorDepth = getParameter errorDepthTag parameters in reportError lex { location = location, hard = true, message = PrettyBlock(0, true, [], [ PrettyString ("Argument for " ^ fnName), PrettyBreak (1, 3), PrettyBlock(0, false, [], [ PrettyString "Required type:", PrettyBreak (1, 0), display (matchType, errorDepth, emptyTypeEnv) ]), PrettyBreak (1, 3), PrettyBlock(0, false, [], [ PrettyString "Argument type:", PrettyBreak (1, 0), display (instanceType, errorDepth, emptyTypeEnv) ]), PrettyBreak (1, 3), unifyTypesErrorReport(lex, emptyTypeEnv, emptyTypeEnv, "unify") error ]), context = SOME (moreInfo ()) } end; (* This is applied to the instance variables if it is polymorphic and bound by a val or fun binding or is a datatype constructor. *) fun applyToInstanceType(polyVars, ValBound, level, typeVarMap, code) = applyToInstance(polyVars, level, typeVarMap, code) | applyToInstanceType(polyVars, Constructor _, level, typeVarMap, code) = applyToInstance(if justForEqualityTypes then [] else polyVars, level, typeVarMap, code) | applyToInstanceType(_, PattBound, level, _, code) = code level | applyToInstanceType(_, Exception, level, _, code) = code level val arg1 = mkLoadArgument 0 (* saves a lot of garbage *) fun addStatus typ = {value=typ, equality=false, printity=false} (* Code-generate an identifier matched to a value. N.B. If the value is a constructor it returns the pair or triple representing the functions on the constructor. *) fun codeVal (Value{access = Global code, class, ...}, level: level, typeVarMap, instance, _, _) = applyToInstanceType(instance, class, level, typeVarMap, fn _ => code) | codeVal (Value{access = Local{addr=ref locAddr, level=ref locLevel}, class, ...}, level, typeVarMap, instance, _, _) = let fun loadVar level = mkLoad (locAddr, level, locLevel) (* Argument or local *) in applyToInstanceType(instance, class, level, typeVarMap, loadVar) end | codeVal (Value{access = Selected{addr, base}, class, ...}, level: level, typeVarMap, instance, _, _) = (* Select from a structure. *) applyToInstanceType(instance, class, level, typeVarMap, fn level => mkInd (addr, codeAccess (base, level))) | codeVal (Value{access = Formal _, ...}, _, _, _, _, _) = raise InternalError "codeVal - Formal" | codeVal (Value{access = Overloaded Print, ...}, _, _, [], lex, _) = (* If this appears in a structure return a null printer function. It has to have the polymorphic form with an extra lambda outside. *) let (* We should have a single entry for the type. *) open DEBUG (* The parameter is the reference used to control the print depth when the value is actually printed. *) val prettyOut = getPrintOutput (LEX.debugParams lex) in mkProc( mkProc( CODETREE.mkEnv ( [ mkNullDec (mkEval( mkConst(toMachineWord prettyOut), [ mkConst(toMachineWord(PrettyString "?")) ]) ) ], arg1 (* Returns its argument. *) ), 1, "print()", [], 0), 1, "print(P)", [], 0) end | codeVal (Value{access = Overloaded Print, ...}, level: level, typeVarMap, [{value=argType, ...}], lex, _) = let (* We should have a single entry for the type. *) open DEBUG (* The parameter is the reference used to control the print depth when the value is actually printed. *) val printDepthFun = getParameter printDepthFunTag (LEX.debugParams lex) and prettyOut = getPrintOutput (LEX.debugParams lex) val nLevel = newLevel level in (* Construct a function that gets the print code, prints it out and returns its argument. *) mkProc( CODETREE.mkEnv ( [ mkNullDec ( mkEval( mkConst(toMachineWord prettyOut), [ mkEval( printerForType(argType, nLevel, typeVarMap), [ mkTuple[arg1, mkEval(mkConst(toMachineWord printDepthFun), [CodeZero])] ]) ]) ) ], arg1 (* Returns its argument. *) ), 1, "print()", getClosure nLevel, 0) end | codeVal (Value{access = Overloaded Print, ...}, _, _, _, _, _) = raise InternalError "Overloaded Print - wrong instance type" | codeVal (Value{access = Overloaded MakeString, ...}, _, _, [], _, _) = (* If this appears in a structure produce a default version. *) mkInlproc( mkProc(mkConst(toMachineWord "?"), 1, "makestring()", [], 0), 1, "makestring(P)", [], 0) | codeVal (Value{access = Overloaded MakeString, ...}, level: level, typeVarMap, [{value=argType, ...}], _, _) = let val nLevel = newLevel level in (* Construct a function that gets the print code and prints it out using "uglyPrint". *) mkProc( mkEval( mkConst(toMachineWord uglyPrint), [ mkEval( printerForType(argType, nLevel, typeVarMap), [ mkTuple[arg1, mkConst(toMachineWord 10000)] ]) ]), 1, "makestring()", getClosure nLevel, 0) end | codeVal (Value{access = Overloaded MakeString, ...}, _, _, _, _, _) = raise InternalError "Overloaded MakeString - wrong instance type" | codeVal (Value{access = Overloaded GetPretty, ...}, level, typeVarMap, [], _, _) = let val nLevel = newLevel level in (* If this appears in a structure return a default function. *) mkProc(printerForType(badType, nLevel, typeVarMap), 1, "getPretty", getClosure nLevel, 0) end | codeVal (Value{access = Overloaded GetPretty, ...}, level: level, typeVarMap, [{value=argType, ...}], _, _) = (* Get the pretty code for the specified argument. *) printerForType(argType, level, typeVarMap) | codeVal (Value{access = Overloaded GetPretty, ...}, _, _, _, _, _) = raise InternalError "Overloaded GetPretty - wrong instance type" | codeVal (Value{access = Overloaded AddPretty, ...}, _, _, [], _, _) = (* If this appears in a structure create a function that raises an exception if run. *) mkProc( mkConst (toMachineWord (fn _ => raise Fail "addPrettyPrint: The argument type was not a simple type construction")), 1, "AddPretty(P)", [], 0) | codeVal (Value{access = Overloaded AddPretty, ...}, level: level, _, [{value=installType, ...}, {value=argPrints, ...}], lex, loc) = let (* "instance" should be (int-> 'a -> 'b -> pretty) -> unit. We need to get the 'a and 'b. This function installs a pretty printer against the type which matches 'b. The type 'a is related to type of 'b as follows: If 'b is a monotype t then 'a is ignored. If 'b is a unary type constructor 'c t then 'a must have type 'c * int -> pretty. If 'b is a binary or higher type constructor e.g. ('c, 'd, 'e) t then 'a must be a tuple of functions of the form ('c * int -> pretty, 'd * int -> pretty, 'e * int -> pretty). When the installed function is called it will be passed the appropriate argument functions which it can call to print the argument types. *) val pretty = mkTypeVar (generalisable, false, false, false); (* Temporary hack. *) (* Find the last type constructor in the chain. We have to install this against the last in the chain because type constructors in different modules may be at different points in the chain. *) (* This does mean that it's not possible to install a pretty printer for a type constructor rather than a datatype. *) fun followTypes (TypeConstruction{constr, args, ...}) = if not (tcIsAbbreviation constr) then SOME(tcIdentifier constr, constr, List.length args) else followTypes (makeEquivalent (constr, args)) | followTypes (TypeVar tv) = ( case tvValue tv of EmptyType => NONE (* Unbound type variable *) | t => followTypes t ) | followTypes _ = NONE; val constrId = followTypes installType val () = case constrId of NONE => () | SOME (_, constr, arity) => let (* Check that the function tuple matches the arguments of the type we're installing for. *) (* Each entry should be a function of type 'a * int -> pretty *) fun mkFn arg = mkFunctionType(mkProductType[arg, TYPESTRUCT.fixedIntType], pretty) (* Create non-unifiable type vars to ensure this is properly polymorphic. *) val typeVars = List.tabulate(arity, fn _ => mkTypeVar (0, false, true, false)) val tupleType = case typeVars of [] => (* No arg so must have unit. *) unitType | [arg] => mkFn arg (* Just a single function. *) | args => mkProductType(List.map mkFn args) val addPPType = mkFunctionType(argPrints, mkFunctionType(installType, pretty)) val testType = mkFunctionType(tupleType, mkFunctionType( mkTypeConstruction(tcName constr, constr, typeVars, [DeclaredAt loc]), pretty)) in checkPPType(addPPType, testType, "addPrettyPrint", lex, loc, fn () => PrettyString "addPrettyPrint element functions must have type 'a * int -> pretty, 'b * int -> pretty, ... with one function for each type parameter") end; (* Only report the error when the function is run. Because addPrettyPrint is contained in the PolyML structure we may compile a reference to a polymorphic version of this for the structure record. It's replaced in the final structure by this version. *) in case constrId of SOME (typeId, _, arity) => let (* We need to transform the user-supplied function into the form required for the reference. The user function has type int -> 'b -> 'a -> pretty where 'b is either "don't care" if this is a monotype, the print function for the base type if it takes a single type argument or a tuple of base type functions if it takes more than one. The reference expects to contain a function of type 'a * int -> pretty for a monotype or a function of the form <'b1, 'b2...> -> 'a * int -> pretty if this is polytype where <...> represents poly-style multiple arguments. *) val printFunction = case arity of 0 => mkProc( mkEval( mkEval( mkEval( mkLoadClosure 0 (* The user-supplied fn *), [mkInd(1, arg1)] (* The depth *)), [CodeZero] (* Ignored args. *)), [mkInd(0, arg1)] (* Value to print *)), 1, "addPP-1", [arg1](* The user-supplied fn *), 0) | arity => let open TypeValue val args = if arity = 1 then [extractPrinter(mkLoadClosure 1)] else [mkTuple(List.tabulate(arity, fn n => extractPrinter(mkLoadClosure(n+1))))] in mkProc( mkProc( mkEval( mkEval( mkEval( mkLoadClosure 0 (* The user-supplied fn *), [mkInd(1, arg1)] (* The depth *)), args (* Base fns. *)), [mkInd(0, arg1)] (* Value to print *)), 1, "addPP-2", mkLoadClosure 0 :: List.tabulate(arity, mkLoadArgument), 0), arity, "addPP-1", [arg1], 0) end val nLevel = newLevel level in (* Generate a function that will set the "print" ref for the type to the argument function. *) mkProc( mkStoreOperation(LoadStoreMLWord{isImmutable=false}, TypeValue.extractPrinter( codeAccess(idAccess typeId, nLevel)), CodeZero, printFunction ), 1, "addPP", getClosure nLevel, 0) end | NONE => mkConst (toMachineWord (fn _ => raise Fail "addPrettyPrint: The argument type was not a simple type construction")) end | codeVal (Value{access = Overloaded AddPretty, ...}, _, _, _, _, _) = raise InternalError "Overloaded AddPretty - wrong instance type" | codeVal (Value{access = Overloaded GetLocation, ...}, _, _, _, _, _) = (* This can't be used a value: It must be called immediately. *) let fun getLoc() = raise Fail "The special function PolyML.sourceLocation cannot be used as a value" in mkConst (toMachineWord getLoc) end | codeVal (value as Value{access = Overloaded _, ...}, level: level, typeVarMap, instance, lex, lineno) = let val nLevel = newLevel level in (* AddOverload, Equal, NotEqual, TypeDep *) mkProc(applyFunction (value, arg1, nLevel, typeVarMap, instance, lex, lineno), 1, "", getClosure nLevel, 0) end (* Some of these have a more efficient way of calling them as functions. *) and applyFunction (value as Value{class=Exception, ...}, argument, level, typeVarMap, instance, lex, lineno) = let (* If we are applying it as a function we cannot be after the exception id, we must be constructing an exception packet. *) (* Get the exception id, put it in the packet with the exception name the argument and, currently, an empty location as the exception location. *) val exIden = codeVal (value, level, typeVarMap, instance, lex, lineno); in mkTuple (exIden :: mkStr (valName value) :: argument :: [mkConst(toMachineWord NoLocation)]) end | applyFunction(value as Value{class=Constructor _, ...}, argument, level, typeVarMap, argVars, lex, lineno) = let (* If this is a value constructor we need to get the construction function and use that. *) fun getConstr level = ValueConstructor.extractInjection(codeVal (value, level, typeVarMap, [], lex, lineno)) val polyConstr = applyToInstance(if justForEqualityTypes then [] else argVars, level, typeVarMap, getConstr) in (* Don't apply this "early". It might be the ref constructor and that must not be applied until run-time. The optimiser should take care of any other cases. *) mkEval (polyConstr, [argument]) end | applyFunction (value as Value{access = Overloaded oper, name = valName, ...}, argument, level, typeVarMap, instance, lex, lineno) = ( case oper of Equal => (* Get the equality function for the type. *) let (* We should have a single entry for the type. *) val argType = case instance of [{value, ...}] => value | _ => raise InternalError "Overload Equal" (* The instance type is a function so we have to get the first argument. *) val code = equalityForType(argType, level, typeVarMap) in mkEval (code, [argument]) end | NotEqual => let (* We should have a single entry for the type. *) val argType = case instance of [{value, ...}] => value | _ => raise InternalError "Overload NotEqual" (* Use the "=" function to provide inequality as well as equality. *) val code = equalityForType(argType, level, typeVarMap) val isEqual = mkEval (code, [argument]) in mkNot isEqual end | TypeDep => let val argType = case instance of [{value, ...}] => value | _ => raise InternalError "Overload TypeDep" val (code, _) = getOverloadInstance(valName, argType, false) in mkEval (code, [argument]) end | AddOverload => (* AddOverload is only intended for use by writers of library modules. It only does limited checking and should be regarded as "unsafe". *) let fun rmvars (TypeVar tv) = rmvars(tvValue tv) | rmvars t = t (* instance should be ('a->'b) -> string -> unit. For overloadings on most functions (e.g. abs and +) we are looking for the 'a, which may be a pair, but in the case of conversion functions we want the 'b. *) val (resultType, argType) = case instance of [{value=alpha, ...}, {value=beta, ...}] => (rmvars alpha, rmvars beta) | _ => (badType, badType) fun followTypes(TypeConstruction{constr as TypeConstrs {identifier = TypeId{idKind = Free _, ...},...}, ...}) = constr | followTypes(TypeConstruction{constr as TypeConstrs {identifier = TypeId{idKind = TypeFn _, ...},...}, args, ...}) = followTypes (makeEquivalent (constr, args)) | followTypes(TypeConstruction{constr = TypeConstrs {identifier = TypeId{idKind = Bound _, ...},...}, ...}) = raise Fail "Cannot install an overload within a structure or functor" | followTypes _ = raise Fail "Invalid type (not a type construction) (addOverload)" fun addOverloading (argCode: codetree) (name: string) = let val typeToUse = if size name > 4 andalso String.substring(name, 0, 4) = "conv" (* For conversion functions it's the result type we're interested in. For everything else it's the argument type. This will be a pair for functions such as "+" and a single argument for "abs". *) then resultType else case argType of LabelledType{recList=[{typeof, ...}, _], ...} => rmvars typeof | argType => argType val tcons = followTypes typeToUse in addOverload(name, tcons, argCode) end (* This function is used if we can't get the codetree at compile time. *) fun addOverloadGeneral (arg: machineWord) = addOverloading(mkConst arg) in (* This is messy but necessary for efficiency. If we simply treat addOverload as a function we would be able to pick up the additional overloading as a pointer to a function. Most overloads are small functions or wrapped calls to RTS functions and so we need to get the inline code for them. *) (* evalue raises an exception if "argument" is not a constant, or more usefully, a global value containing a constant and possibly a piece of codetree to inline. *) case evalue(argument) of SOME _ => mkConst (toMachineWord (addOverloading argument)) | NONE => mkEval (mkConst (toMachineWord addOverloadGeneral), [argument]) end | GetLocation => (* Return the current location. *) mkConst(toMachineWord lineno) | _ => (* Print, MakeString, InstallPP *) (* Just call as functions. *) (* not early *) mkEval (codeVal (value, level, typeVarMap, instance, lex, lineno), [argument]) ) (* overloaded *) | applyFunction (value, argument, level, typeVarMap, instance, lex, lineno) = mkEval (codeVal (value, level, typeVarMap, instance, lex, lineno), [argument]) (* end applyFunction *) (* If the exception is being used as a value we want an exception packet or a function to make a packet. If it is a nullary constructor make an exception packet now, otherwise generate a function to construct an exception packet. *) fun codeExFunction (value, level, typeVarMap, instance, lex, lineno) = case getFnArgType(valTypeOf value) of (* N.B. Not "instance" *) NONE => applyFunction (value, CodeZero, level, typeVarMap, List.map addStatus instance, lex, lineno) | SOME _ => let val nLevel = newLevel level in mkProc (applyFunction (value, arg1, nLevel, typeVarMap, List.map addStatus instance, lex, lineno), 1, "", getClosure nLevel, 0) end (* Operations to compile code from the representation of a constructor. *) (* Code to test whether a value matches a constructor. This must be applied to any polymorphic variables in the instance but the result is always bool so we don't create a new function if the result is also polymorphic. It is just possible to have a resulting polytype here (N.B. that's different from having a parametric type) if we have a val binding. e.g. val SOME x = SOME nil. In that case we can choose an arbitrary type for the test and have to parameterise the result. *) fun makeGuard (value as Value{class=Constructor _, ...}, argVars, testing, level, typeVarMap) = let fun tester level = ValueConstructor.extractTest(codeVal (value, level, typeVarMap, [], nullLex, location nullLex)) val testCode = applyToInstance(if justForEqualityTypes then [] else List.map addStatus argVars, level, typeVarMap, tester) in mkEval(testCode, [testing]) end | makeGuard (value as Value{class=Exception, ...}, _, testing, level, typeVarMap) = (* Should only be an exception. Get the value of the exception identifier and compare with the identifier in the exception packet. *) - mkEqualWord (mkInd (0, testing), + mkEqualPointerOrWord (mkInd (0, testing), codeVal (value, level, typeVarMap, [], nullLex, location nullLex)) | makeGuard _ = raise InternalError "makeGuard" (* Code to invert a constructor. i.e. return the value originally used as the argument. Apply to any polymorphic variables and construct a result. *) fun makeInverse(value as Value{class=Constructor{nullary=false, ...}, ...}, argVars, arg, level, typeVarMap): codetree = let fun getInverse level = ValueConstructor.extractProjection(codeVal (value, level, typeVarMap, [], nullLex, location nullLex)) val loadCode = applyToInstance(if justForEqualityTypes then [] else List.map addStatus argVars, level, typeVarMap, getInverse) in mkEval(loadCode, [arg]) end | makeInverse(Value{class=Constructor{nullary=true, ...}, ...}, _, _, _, _): codetree = (* makeInverse is called even on nullary constructors. Return zero to keep the optimiser happy. *) CodeZero | makeInverse (Value{class=Exception, ...}, _, arg, _, _) = (* Exceptions. - Get the parameter from third word *) (* We have to use a VarField here even though this field is present in every exception. The format of the value that is returned depends on the exception id. *) mkVarField (2,arg) | makeInverse _ = raise InternalError "makeInverse" (* Work out the polymorphism and the mapping between the formal type variables and the actual types. Because flexible records may introduce extra polymorphism we can only do this once we've frozen them. e.g. fun f x = #1 x + #2 x may be monomorphic or polymorphic depending on what it's subsequently applied to. *) (* Using unification here isn't ideal. We have to put the equality attribute back on to abstypes in case the unification requires it. There may be other situations where things don't work properly. *) fun getPolymorphism (Value{ typeOf, access, name, ...}, expType, typeVarMap) = let val (t, polyVars) = case access of Overloaded TypeDep => let val (t, polyVars) = generaliseOverload(typeOf, List.map #1 (getOverloads name), false) in (t, List.map (fn t => {value=t, equality=false, printity=false}) polyVars) end | _ => generaliseWithMap(typeOf, TypeVarMap.mapTypeVars typeVarMap) (* Ignore the result. There are circumstances in which we can get a unification error as the result of failing to find a fixed record type where the possible records we could find have non-unifiable types. See Tests/Fail/Test072.ML *) val _ = unifyTypes(t, expType) in polyVars end (* Convert a literal constant. We can only do this once any overloading has been resolved. *) fun getLiteralValue(converter, literal, instance, error): machineWord option = let val (conv, name) = getOverloadInstance(valName converter, instance, true) in SOME(RunCall.unsafeCast(valOf(evalue conv)) literal) handle Match => NONE (* Overload error *) | Conversion s => ( error("Conversion exception ("^s^") raised while converting " ^ literal ^ " to " ^ name); NONE ) | Overflow => ( error ("Overflow exception raised while converting " ^ literal ^ " to " ^ name); NONE ) | Thread.Thread.Interrupt => raise Thread.Thread.Interrupt | _ => ( error ("Exception raised while converting " ^ literal ^ " to " ^ name); NONE ) end (* Types that can be shared. *) structure Sharing = struct type lexan = lexan type codetree = codetree type types = types type values = values type structVals = structVals type functors = functors type valAccess = valAccess type typeConstrs = typeConstrs type typeConstrSet = typeConstrSet type signatures = signatures type fixStatus = fixStatus type univTable = univTable type pretty = pretty type locationProp = locationProp type typeId = typeId type typeVarForm = typeVarForm type typeVarMap = typeVarMap type level = level type machineWord = machineWord end end (* body of VALUEOPS *);