From 1abcc8da955eefdbd578580d9410f5cb9ef5df60 Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Fri, 8 Jul 2022 17:45:34 -0700 Subject: [PATCH 01/13] Add StructuredEncryption smithy model and stubbed APIs --- .gitignore | 16 + .gitmodules | 6 + README.md | 29 +- libraries | 1 + polymorph | 1 + src/StandardLibrary/StandardLibrary.dfy | 406 ++++++++++++++++++ src/StandardLibrary/UInt.dfy | 332 ++++++++++++++ ...sCryptographyStructuredEncryptionTypes.dfy | 129 ++++++ .../model/StructuredEncryption.smithy | 136 ++++++ src/StructuredEncryption/model/traits.smithy | 8 + .../Test/AwsStructuredEncryptionTest.csproj | 80 ++++ .../src/StructuredEncryptionClient.dfy | 100 +++++ .../src/StructuredEncryptionDafny.csproj | 30 ++ .../test/HappyCaseTests.dfy | 102 +++++ src/Util/UTF8.dfy | 154 +++++++ 15 files changed, 1525 insertions(+), 5 deletions(-) create mode 100644 .gitignore create mode 100644 .gitmodules create mode 160000 libraries create mode 160000 polymorph create mode 100644 src/StandardLibrary/StandardLibrary.dfy create mode 100644 src/StandardLibrary/UInt.dfy create mode 100644 src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy create mode 100644 src/StructuredEncryption/model/StructuredEncryption.smithy create mode 100644 src/StructuredEncryption/model/traits.smithy create mode 100644 src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj create mode 100644 src/StructuredEncryption/src/StructuredEncryptionClient.dfy create mode 100644 src/StructuredEncryption/src/StructuredEncryptionDafny.csproj create mode 100644 src/StructuredEncryption/test/HappyCaseTests.dfy create mode 100644 src/Util/UTF8.dfy diff --git a/.gitignore b/.gitignore new file mode 100644 index 000000000..443fb2436 --- /dev/null +++ b/.gitignore @@ -0,0 +1,16 @@ +.idea/* +.vscode/* +*.sln* +**/.DS_Store +**/obj +**/bin +build/* +test/**/Output/* +/package-lock.json +/node_modules + +# Duvet output +specification_compliance_report.html + +# VS Code local history +/.history diff --git a/.gitmodules b/.gitmodules new file mode 100644 index 000000000..27d8705bf --- /dev/null +++ b/.gitmodules @@ -0,0 +1,6 @@ +[submodule "libraries"] + path = libraries + url = git@github.com:dafny-lang/libraries.git +[submodule "polymorph"] + path = polymorph + url = git@github.com:awslabs/polymorph.git diff --git a/README.md b/README.md index 847260ca5..cb69d7eb6 100644 --- a/README.md +++ b/README.md @@ -1,11 +1,30 @@ -## My Project +## DynamoDB Encryption Client for Dafny -TODO: Fill this README out! +TODO: Edit your repository description on GitHub -Be sure to: +### Rebuild the model -* Change the title in this README -* Edit your repository description on GitHub +Currently only smithy -> dafny is working. Need to get smithy -> .NET with union support. + +Run the following: +``` +cd polymorph/smithy-dotnet +export MODEL_ROOT=/aws-dynamodb-encryption-dafny/src/StructuredEncryption/model + +./gradlew run --args="\ + -m $MODEL_ROOT \ + -namespace aws.cryptography.structuredEncryption \ + -d $MODEL_ROOT" +``` + +### Run the tests + +Assuming you are on a Mac: + +``` +cd src/StructuredEncryption/runtimes/net +dotnet test -f netcoreapp3.1 Test +``` ## Security diff --git a/libraries b/libraries new file mode 160000 index 000000000..735839fa6 --- /dev/null +++ b/libraries @@ -0,0 +1 @@ +Subproject commit 735839fa6e79ae60cc90fb529bf0cdb367e26b00 diff --git a/polymorph b/polymorph new file mode 160000 index 000000000..bff5c35aa --- /dev/null +++ b/polymorph @@ -0,0 +1 @@ +Subproject commit bff5c35aa3323ebe9b7ff38abcb27b01486d992e diff --git a/src/StandardLibrary/StandardLibrary.dfy b/src/StandardLibrary/StandardLibrary.dfy new file mode 100644 index 000000000..a5dbf3abb --- /dev/null +++ b/src/StandardLibrary/StandardLibrary.dfy @@ -0,0 +1,406 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 + +include "../../libraries/src/Wrappers.dfy" +include "UInt.dfy" + +module StandardLibrary { + import opened Wrappers + import opened U = UInt + + function method {:tailrecursion} Join(ss: seq>, joiner: seq): (s: seq) + requires 0 < |ss| + { + if |ss| == 1 then ss[0] else ss[0] + joiner + Join(ss[1..], joiner) + } + + function method {:tailrecursion} Split(s: seq, delim: T): (res: seq>) + ensures delim !in s ==> res == [s] + ensures s == [] ==> res == [[]] + ensures 0 < |res| + ensures forall i :: 0 <= i < |res| ==> delim !in res[i] + ensures Join(res, [delim]) == s + decreases |s| + { + var i := FindIndexMatching(s, delim, 0); + if i.Some? then [s[..i.value]] + Split(s[(i.value + 1)..], delim) else [s] + } + + lemma WillSplitOnDelim(s: seq, delim: T, prefix: seq) + requires |prefix| < |s| + requires forall i :: 0 <= i < |prefix| ==> prefix[i] == s[i] + requires delim !in prefix && s[|prefix|] == delim + ensures Split(s, delim) == [prefix] + Split(s[|prefix| + 1..], delim) + { + calc { + Split(s, delim); + == + var i := FindIndexMatching(s, delim, 0); + if i.Some? then [s[..i.value]] + Split(s[i.value + 1..], delim) else [s]; + == { FindIndexMatchingLocatesElem(s, delim, 0, |prefix|); assert FindIndexMatching(s, delim, 0).Some?; } + [s[..|prefix|]] + Split(s[|prefix| + 1..], delim); + == { assert s[..|prefix|] == prefix; } + [prefix] + Split(s[|prefix| + 1..], delim); + } + } + + lemma WillNotSplitWithOutDelim(s: seq, delim: T) + requires delim !in s + ensures Split(s, delim) == [s] + { + calc { + Split(s, delim); + == + var i := FindIndexMatching(s, delim, 0); + if i.Some? then [s[..i.value]] + Split(s[i.value+1..], delim) else [s]; + == { FindIndexMatchingLocatesElem(s, delim, 0, |s|); } + [s]; + } + } + + lemma FindIndexMatchingLocatesElem(s: seq, c: T, start: nat, elemIndex: nat) + requires start <= elemIndex <= |s| + requires forall i :: start <= i < elemIndex ==> s[i] != c + requires elemIndex == |s| || s[elemIndex] == c + ensures FindIndexMatching(s, c, start) == if elemIndex == |s| then None else Some(elemIndex) + decreases elemIndex - start + {} + + function method FindIndexMatching(s: seq, c: T, i: nat): (index: Option) + requires i <= |s| + ensures index.Some? ==> i <= index.value < |s| && s[index.value] == c && c !in s[i..index.value] + ensures index.None? ==> c !in s[i..] + decreases |s| - i + { + FindIndex(s, x => x == c, i) + } + + function method {:tailrecursion} FindIndex(s: seq, f: T -> bool, i: nat): (index: Option) + requires i <= |s| + ensures index.Some? ==> i <= index.value < |s| && f(s[index.value]) && (forall j :: i <= j < index.value ==> !f(s[j])) + ensures index.None? ==> forall j :: i <= j < |s| ==> !f(s[j]) + decreases |s| - i + { + if i == |s| then None + else if f(s[i]) then Some(i) + else FindIndex(s, f, i + 1) + } + + function method {:tailrecursion} Filter(s: seq, f: T -> bool): (res: seq) + ensures forall i :: 0 <= i < |s| && f(s[i]) ==> s[i] in res + ensures forall i :: 0 <= i < |res| ==> res[i] in s && f(res[i]) + ensures |res| <= |s| + { + if |s| == 0 then [] + else if f(s[0]) then [s[0]] + Filter(s[1..], f) + else Filter(s[1..], f) + } + + lemma FilterIsDistributive(s: seq, s': seq, f: T -> bool) + ensures Filter(s + s', f) == Filter(s, f) + Filter(s', f) + { + if s == [] { + assert s + s' == s'; + } else { + var S := s + s'; + var s1 := s[1..]; + calc { + Filter(S, f); + == // def. Filter + if f(S[0]) then [S[0]] + Filter(S[1..], f) else Filter(S[1..], f); + == { assert S[0] == s[0] && S[1..] == s1 + s'; } + if f(s[0]) then [s[0]] + Filter(s1 + s', f) else Filter(s1 + s', f); + == { FilterIsDistributive(s1, s', f); } + if f(s[0]) then [s[0]] + (Filter(s1, f) + Filter(s', f)) else Filter(s1, f) + Filter(s', f); + == // associativity of + + if f(s[0]) then ([s[0]] + Filter(s1, f)) + Filter(s', f) else Filter(s1, f) + Filter(s', f); + == // distribute + over if-then-else + (if f(s[0]) then [s[0]] + Filter(s1, f) else Filter(s1, f)) + Filter(s', f); + == // def. Filter + Filter(s, f) + Filter(s', f); + } + } + } + + function method Min(a: int, b: int): int { + if a < b then a else b + } + + function method Fill(value: T, n: nat): seq + ensures |Fill(value, n)| == n + ensures forall i :: 0 <= i < n ==> Fill(value, n)[i] == value + { + seq(n, _ => value) + } + + method SeqToArray(s: seq) returns (a: array) + // "Fresh" expressions require editing memory + ensures fresh(a) + ensures a.Length == |s| + ensures forall i :: 0 <= i < |s| ==> a[i] == s[i] + { + a := new T[|s|](i requires 0 <= i < |s| => s[i]); + } + + lemma SeqPartsMakeWhole(s: seq, i: nat) + requires 0 <= i <= |s| + ensures s[..i] + s[i..] == s + {} + + /* + * Lexicographic comparison of sequences. + * + * Given two sequences `a` and `b` and a strict (that is, irreflexive) + * ordering `less` on the elements of these sequences, determine whether or not + * `a` is lexicographically "less than or equal to" `b`. + * + * `a` is lexicographically "less than or equal to" `b` holds iff + * there exists a `k` such that + * - the first `k` elements of `a` and `b` are the same + * - either: + * -- `a` has length `k` (that is, `a` is a prefix of `b`) + * -- `a[k]` is strictly less (using `less`) than `b[k]` + */ + + predicate method LexicographicLessOrEqual(a: seq, b: seq, less: (T, T) -> bool) { + exists k :: 0 <= k <= |a| && LexicographicLessOrEqualAux(a, b, less, k) + } + + predicate method LexicographicLessOrEqualAux(a: seq, b: seq, less: (T, T) -> bool, lengthOfCommonPrefix: nat) + requires 0 <= lengthOfCommonPrefix <= |a| + { + lengthOfCommonPrefix <= |b| + && (forall i :: 0 <= i < lengthOfCommonPrefix ==> a[i] == b[i]) + && (lengthOfCommonPrefix == |a| || (lengthOfCommonPrefix < |b| && less(a[lengthOfCommonPrefix], b[lengthOfCommonPrefix]))) + } + + /* + * In order for the lexicographic ordering above to have the expected properties, the + * relation "less" must be trichotomous and transitive. + * + * For an ordering `less` to be _trichotomous_ means that for any two `x` and `y`, + * EXACTLY one of the following three conditions holds: + * - less(x, y) + * - x == y + * - less(y, x) + * Note that being trichotomous implies being irreflexive. + */ + + predicate Trichotomous(less: (T, T) -> bool) { + (forall x, y :: less(x, y) || x == y || less(y, x)) && // at least one of the three + (forall x, y :: less(x, y) && less(y, x) ==> false) && // not both of the less's + (forall x, y :: less(x, y) ==> x != y) // not a less and the equality + } + + predicate Transitive(R: (T, T) -> bool) { + forall x, y, z :: R(x, y) && R(y, z) ==> R(x, z) + } + + /* + * Here is an example relation and a lemma that says the relation is appropriate for use in + * lexicographic orderings. + */ + + lemma UInt8LessIsTrichotomousTransitive() + ensures Trichotomous(UInt8Less) + ensures Transitive(UInt8Less) + { + } + + /* + * As the following lemmas show, the lexicographic ordering is reflexive, antisymmetric, transitive, and total. + * The proofs are a bit pedantic and include steps that can be automated. + */ + + lemma LexIsReflexive(a: seq, less: (T, T) -> bool) + ensures LexicographicLessOrEqual(a, a, less) + { + assert LexicographicLessOrEqualAux(a, a, less, |a|); + } + + lemma LexIsAntisymmetric(a: seq, b: seq, less: (T, T) -> bool) + requires Trich: Trichotomous(less) + requires LexicographicLessOrEqual(a, b, less) + requires LexicographicLessOrEqual(b, a, less) + ensures a == b + { + assert LessIrreflexive: forall x,y :: less(x, y) ==> x != y by { + reveal Trich; + } + assert ASymmetric: forall x,y :: less(x, y) && less(y, x) ==> false by { + reveal Trich; + } + var k0 :| 0 <= k0 <= |a| && LexicographicLessOrEqualAux(a, b, less, k0); + var k1 :| 0 <= k1 <= |b| && LexicographicLessOrEqualAux(b, a, less, k1); + var max := if k0 < k1 then k1 else k0; + assert max <= |a| && max <= |b|; + assert SameUntilMax: forall i :: 0 <= i < max ==> a[i] == b[i]; + assert AA: k0 == |a| || (k0 < |b| && less(a[k0], b[k0])); + assert BB: k1 == |b| || (k1 < |a| && less(b[k1], a[k1])); + calc { + true; + ==> { reveal AA, BB; } + (k0 == |a| || (k0 < |b| && less(a[k0], b[k0]))) && (k1 == |b| || (k1 < |a| && less(b[k1], a[k1]))); + == // distribute && and || + (k0 == |a| && k1 == |b|) || + (k0 == |a| && k1 < |a| && less(b[k1], a[k1])) || + (k0 < |b| && less(a[k0], b[k0]) && k1 == |b|) || + (k0 < |b| && less(a[k0], b[k0]) && k1 < |a| && less(b[k1], a[k1])); + == { reveal LessIrreflexive, SameUntilMax; } + (k0 == |a| && k1 == |b|) || + (k0 < |b| && less(a[k0], b[k0]) && k1 < |a| && less(b[k1], a[k1])); + ==> { reveal LessIrreflexive, SameUntilMax; assert max <= k0 && max <= k1; } + (k0 == |a| && k1 == |b|) || + (k0 < |b| && less(a[k0], b[k0]) && k1 < |a| && less(b[k1], a[k1]) && k0 == k1 == max); + == { reveal ASymmetric; } + k0 == |a| && k1 == |b|; + ==> { assert |a| == k0 <= max && |b| == k1 <= max ==> k0 == k1; } + max == |a| == |b|; + ==> { reveal SameUntilMax; } + a == b; + } + } + + lemma LexIsTransitive(a: seq, b: seq, c: seq, less: (T, T) -> bool) + requires Transitive(less) + requires LexicographicLessOrEqual(a, b, less) + requires LexicographicLessOrEqual(b, c, less) + ensures LexicographicLessOrEqual(a, c, less) + { + var k0 :| 0 <= k0 <= |a| && LexicographicLessOrEqualAux(a, b, less, k0); + var k1 :| 0 <= k1 <= |b| && LexicographicLessOrEqualAux(b, c, less, k1); + var k := if k0 < k1 then k0 else k1; + assert LexicographicLessOrEqualAux(a, c, less, k); + } + + lemma LexIsTotal(a: seq, b: seq, less: (T, T) -> bool) + requires Trich: Trichotomous(less) + ensures LexicographicLessOrEqual(a, b, less) || LexicographicLessOrEqual(b, a, less) + { + var m := 0; + while m < |a| && m < |b| && a[m] == b[m] + invariant m <= |a| && m <= |b| + invariant forall i :: 0 <= i < m ==> a[i] == b[i] + { + m := m + 1; + } + // m is the length of the common prefix of a and b + if m == |a| == |b| { + assert a == b; + LexIsReflexive(a, less); + } else if m == |a| < |b| { + assert LexicographicLessOrEqualAux(a, b, less, m); + } else if m == |b| < |a| { + assert LexicographicLessOrEqualAux(b, a, less, m); + } else { + assert m < |a| && m < |b|; + reveal Trich; + if + case less(a[m], b[m]) => + assert LexicographicLessOrEqualAux(a, b, less, m); + case less(b[m], a[m]) => + assert LexicographicLessOrEqualAux(b, a, less, m); + } + } + + /* + * SetToOrderedSequence(s, less) takes a set of T-strings and returns them as a sequence, + * ordered by the lexicographic ordering whose underlying irreflexive ordering is "less". + * The function is compilable, but will not exhibit enviable performance. + */ + + function method {:tailrecursion} SetToOrderedSequence(s: set>, less: (T, T) -> bool): (q: seq>) + requires Trichotomous(less) && Transitive(less) + ensures |s| == |q| + ensures forall i :: 0 <= i < |q| ==> q[i] in s + ensures forall k :: k in s ==> k in q + ensures forall i :: 0 < i < |q| ==> LexicographicLessOrEqual(q[i-1], q[i], less) + ensures forall i, j | 0 <= i < j < |q| :: q[i] != q[j]; + { + if s == {} then + [] + else + // In preparation for the assign-such-that statement below, we'll need to + // prove that a minimum exists and that it is unique. + // The following lemma shows existence: + ThereIsAMinimum(s, less); + // The following assertion shows uniqueness: + assert forall a, b :: IsMinimum(a, s, less) && IsMinimum(b, s, less) ==> a == b by { + // The proof of the assertion is the following forall statement. + // But why did we even bother to write the assert-by instead of + // just writing this forall statement in the first place? Because + // we are in an expression context and a forall statement cannot start + // an expression (because the "forall" makes the parser think that + // a forall expression is coming). + forall a, b | IsMinimum(a, s, less) && IsMinimum(b, s, less) { + // For the given a and b, the proof is settled by calling the following lemma: + MinimumIsUnique(a, b, s, less); + } + } + // The "a in s" in the following line follows from IsMinimum(a, s), so it + // is logically redundant. However, it is needed to convince the compiler + // that the assign-such-that statement is compilable. + var a :| a in s && IsMinimum(a, s, less); + [a] + SetToOrderedSequence(s - {a}, less) + } + + predicate method IsMinimum(a: seq, s: set>, less: (T, T) -> bool) { + a in s && + forall z :: z in s ==> LexicographicLessOrEqual(a, z, less) + } + + lemma ThereIsAMinimum(s: set>, less: (T, T) -> bool) + requires s != {} + requires Trichotomous(less) && Transitive(less) + ensures exists a :: IsMinimum(a, s, less) + { + var a := FindMinimum(s, less); + } + + lemma MinimumIsUnique(a: seq, b: seq, s: set>, less: (T, T) -> bool) + requires IsMinimum(a, s, less) && IsMinimum(b, s, less) + requires Trichotomous(less) + ensures a == b + { + LexIsAntisymmetric(a, b, less); + } + + lemma FindMinimum(s: set>, less: (T, T) -> bool) returns (a: seq) + requires s != {} + requires Trichotomous(less) && Transitive(less) + ensures IsMinimum(a, s, less) + { + a :| a in s; + if s == {a} { + LexIsReflexive(a, less); + } else { + var s' := s - {a}; + assert forall x :: x in s <==> x == a || x in s'; + var a' := FindMinimum(s', less); + if LexicographicLessOrEqual(a', a, less) { + a := a'; + } else { + assert LexicographicLessOrEqual(a, a', less) by { + LexIsTotal(a, a', less); + } + forall z | z in s + ensures LexicographicLessOrEqual(a, z, less) + { + if z == a { + LexIsReflexive(a, less); + } else { + calc { + true; + == // z in s && z != a + z in s'; + ==> // by postcondition of FindMinim(s') above + LexicographicLessOrEqual(a', z, less); + ==> { LexIsTransitive(a, a', z, less); } + LexicographicLessOrEqual(a, z, less); + } + } + } + } + } + } +} diff --git a/src/StandardLibrary/UInt.dfy b/src/StandardLibrary/UInt.dfy new file mode 100644 index 000000000..7f4613147 --- /dev/null +++ b/src/StandardLibrary/UInt.dfy @@ -0,0 +1,332 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 + +module StandardLibrary.UInt { + + // TODO: Depend on types defined in dafny-lang/libraries instead + newtype uint8 = x | 0 <= x < 0x100 + newtype uint16 = x | 0 <= x < 0x1_0000 + newtype uint32 = x | 0 <= x < 0x1_0000_0000 + newtype uint64 = x | 0 <= x < 0x1_0000_0000_0000_0000 + + newtype int32 = x | -0x8000_0000 <= x < 0x8000_0000 + newtype int64 = x | -0x8000_0000_0000_0000 <= x < 0x8000_0000_0000_0000 + newtype posInt64 = x | 0 < x < 0x8000_0000_0000_0000 witness 1 + + const UINT16_LIMIT := 0x1_0000 + const UINT32_LIMIT := 0x1_0000_0000 + const UINT64_LIMIT := 0x1_0000_0000_0000_0000 + const INT32_MAX_LIMIT := 0x8000_0000 + const INT64_MAX_LIMIT := 0x8000_0000_0000_0000 + + predicate method UInt8Less(a: uint8, b: uint8) { a < b } + + predicate method HasUint16Len(s: seq) { + |s| < UINT16_LIMIT + } + + type seq16 = s: seq | HasUint16Len(s) + type Uint8Seq16 = seq16 + + predicate method HasUint32Len(s: seq) { + |s| < UINT32_LIMIT + } + + type seq32 = s: seq | HasUint32Len(s) + type Uint8Seq32 = seq32 + + predicate method HasUint64Len(s: seq) { + |s| < UINT64_LIMIT + } + + type seq64 = s: seq | HasUint64Len(s) + type Uint8Seq64 = seq64 + + function method UInt16ToSeq(x: uint16): (ret: seq) + ensures |ret| == 2 + ensures 0x100 * ret[0] as uint16 + ret[1] as uint16 == x + { + var b0 := (x / 0x100) as uint8; + var b1 := (x % 0x100) as uint8; + [b0, b1] + } + + function method SeqToUInt16(s: seq): (x: uint16) + requires |s| == 2 + ensures UInt16ToSeq(x) == s + ensures x >= 0 + { + var x0 := s[0] as uint16 * 0x100; + x0 + s[1] as uint16 + } + + lemma UInt16SeqSerializeDeserialize(x: uint16) + ensures SeqToUInt16(UInt16ToSeq(x)) == x + {} + + lemma UInt16SeqDeserializeSerialize(s: seq) + requires |s| == 2 + ensures UInt16ToSeq(SeqToUInt16(s)) == s + {} + + function method UInt32ToSeq(x: uint32): (ret: seq) + ensures |ret| == 4 + ensures 0x100_0000 * ret[0] as uint32 + 0x1_0000 * ret[1] as uint32 + 0x100 * ret[2] as uint32 + ret[3] as uint32 == x + { + var b0 := (x / 0x100_0000) as uint8; + var x0 := x - (b0 as uint32 * 0x100_0000); + + var b1 := (x0 / 0x1_0000) as uint8; + var x1 := x0 - (b1 as uint32 * 0x1_0000); + + var b2 := (x1 / 0x100) as uint8; + var b3 := (x1 % 0x100) as uint8; + [b0, b1, b2, b3] + } + + function method SeqToUInt32(s: seq): (x: uint32) + requires |s| == 4 + ensures UInt32ToSeq(x) == s + { + var x0 := s[0] as uint32 * 0x100_0000; + var x1 := x0 + s[1] as uint32 * 0x1_0000; + var x2 := x1 + s[2] as uint32 * 0x100; + x2 + s[3] as uint32 + } + + lemma UInt32SeqSerializeDeserialize(x: uint32) + ensures SeqToUInt32(UInt32ToSeq(x)) == x + {} + + lemma UInt32SeqDeserializeSerialize(s: seq) + requires |s| == 4 + ensures UInt32ToSeq(SeqToUInt32(s)) == s + {} + + function method UInt64ToSeq(x: uint64): (ret: seq) + ensures |ret| == 8 + ensures 0x100_0000_0000_0000 * ret[0] as uint64 + 0x1_0000_0000_0000 * ret[1] as uint64 + + 0x100_0000_0000 * ret[2] as uint64 + 0x1_0000_0000 * ret[3] as uint64 + 0x100_0000 * ret[4] as uint64 + + 0x1_0000 * ret[5] as uint64 + 0x100 * ret[6] as uint64 + ret[7] as uint64 == x + { + var b0 := (x / 0x100_0000_0000_0000) as uint8; + var x0 := x - (b0 as uint64 * 0x100_0000_0000_0000); + + var b1 := (x0 / 0x1_0000_0000_0000) as uint8; + var x1 := x0 - (b1 as uint64 * 0x1_0000_0000_0000); + + var b2 := (x1 / 0x100_0000_0000) as uint8; + var x2 := x1 - (b2 as uint64 * 0x100_0000_0000); + + var b3 := (x2 / 0x1_0000_0000) as uint8; + var x3 := x2 - (b3 as uint64 * 0x1_0000_0000); + + var b4 := (x3 / 0x100_0000) as uint8; + var x4 := x3 - (b4 as uint64 * 0x100_0000); + + var b5 := (x4 / 0x1_0000) as uint8; + var x5 := x4 - (b5 as uint64 * 0x1_0000); + + var b6 := (x5 / 0x100) as uint8; + var b7 := (x5 % 0x100) as uint8; + [b0, b1, b2, b3, b4, b5, b6, b7] + } + + function method SeqToUInt64(s: seq): (x: uint64) + requires |s| == 8 + ensures UInt64ToSeq(x) == s + { + var x0 := s[0] as uint64 * 0x100_0000_0000_0000; + var x1 := x0 + s[1] as uint64 * 0x1_0000_0000_0000; + var x2 := x1 + s[2] as uint64 * 0x100_0000_0000; + var x3 := x2 + s[3] as uint64 * 0x1_0000_0000; + var x4 := x3 + s[4] as uint64 * 0x100_0000; + var x5 := x4 + s[5] as uint64 * 0x1_0000; + var x6 := x5 + s[6] as uint64 * 0x100; + var x := x6 + s[7] as uint64; + UInt64SeqSerialize(x, s); + x + } + + lemma UInt64SeqSerialize(x: uint64, s: seq) + requires |s| == 8 + requires 0x100_0000_0000_0000 * s[0] as uint64 + + 0x1_0000_0000_0000 * s[1] as uint64 + + 0x100_0000_0000 * s[2] as uint64 + + 0x1_0000_0000 * s[3] as uint64 + + 0x100_0000 * s[4] as uint64 + + 0x1_0000 * s[5] as uint64 + + 0x100 * s[6] as uint64 + + s[7] as uint64 == x + ensures UInt64ToSeq(x) == s + { + calc { + UInt64ToSeq(x); + == + UInt64ToSeq(s[0] as uint64 * 0x100_0000_0000_0000 + + s[1] as uint64 * 0x1_0000_0000_0000 + + s[2] as uint64 * 0x100_0000_0000 + + s[3] as uint64 * 0x1_0000_0000 + + s[4] as uint64 * 0x100_0000 + + s[5] as uint64 * 0x1_0000 + + s[6] as uint64 * 0x100 + + s[7] as uint64); + == + var b0 := ((s[0] as uint64 * 0x100_0000_0000_0000 + + s[1] as uint64 * 0x1_0000_0000_0000 + + s[2] as uint64 * 0x100_0000_0000 + + s[3] as uint64 * 0x1_0000_0000 + + s[4] as uint64 * 0x100_0000 + + s[5] as uint64 * 0x1_0000 + + s[6] as uint64 * 0x100 + + s[7] as uint64) / 0x100_0000_0000_0000) as uint8; + assert b0 == s[0]; + var x0 := (s[0] as uint64 * 0x100_0000_0000_0000 + + s[1] as uint64 * 0x1_0000_0000_0000 + + s[2] as uint64 * 0x100_0000_0000 + + s[3] as uint64 * 0x1_0000_0000 + + s[4] as uint64 * 0x100_0000 + + s[5] as uint64 * 0x1_0000 + + s[6] as uint64 * 0x100 + + s[7] as uint64) - (b0 as uint64 * 0x100_0000_0000_0000); + assert x0 == (s[1] as uint64 * 0x1_0000_0000_0000 + + s[2] as uint64 * 0x100_0000_0000 + + s[3] as uint64 * 0x1_0000_0000 + + s[4] as uint64 * 0x100_0000 + + s[5] as uint64 * 0x1_0000 + + s[6] as uint64 * 0x100 + + s[7] as uint64); + + var b1 := (x0 / 0x1_0000_0000_0000) as uint8; + assert b1 == s[1]; + var x1 := x0 - (b1 as uint64 * 0x1_0000_0000_0000); + assert x1 == (s[2] as uint64 * 0x100_0000_0000 + + s[3] as uint64 * 0x1_0000_0000 + + s[4] as uint64 * 0x100_0000 + + s[5] as uint64 * 0x1_0000 + + s[6] as uint64 * 0x100 + + s[7] as uint64); + + var b2 := (x1 / 0x100_0000_0000) as uint8; + assert b2 == s[2]; + var x2 := x1 - (b2 as uint64 * 0x100_0000_0000); + assert x2 == (s[3] as uint64 * 0x1_0000_0000 + + s[4] as uint64 * 0x100_0000 + + s[5] as uint64 * 0x1_0000 + + s[6] as uint64 * 0x100 + + s[7] as uint64); + + var b3 := (x2 / 0x1_0000_0000) as uint8; + assert b3 == s[3]; + var x3 := x2 - (b3 as uint64 * 0x1_0000_0000); + assert x3 == (s[4] as uint64 * 0x100_0000 + + s[5] as uint64 * 0x1_0000 + + s[6] as uint64 * 0x100 + + s[7] as uint64); + + var b4 := (x3 / 0x100_0000) as uint8; + assert b4 == s[4]; + var x4 := x3 - (b4 as uint64 * 0x100_0000); + assert x4 == (s[5] as uint64 * 0x1_0000 + + s[6] as uint64 * 0x100 + + s[7] as uint64); + + var b5 := (x4 / 0x1_0000) as uint8; + assert b5 == s[5]; + var x5 := x4 - (b5 as uint64 * 0x1_0000); + assert x5 == (s[6] as uint64 * 0x100 + s[7] as uint64); + + var b6 := (x5 / 0x100) as uint8; + assert b6 == s[6]; + var b7 := (x5 % 0x100) as uint8; + assert b7 == s[7]; + [b0, b1, b2, b3, b4, b5, b6, b7]; + == + [s[0], s[1], s[2], s[3], s[4], s[5], s[6], s[7]]; + == + s; + } + } + + lemma UInt64SeqSerializeDeserialize(x: uint64) + ensures SeqToUInt64(UInt64ToSeq(x)) == x + {} + + lemma UInt64SeqDeserializeSerialize(s: seq) + requires |s| == 8 + ensures UInt64ToSeq(SeqToUInt64(s)) == s + {} + + function SeqToNat(s: seq): nat { + if s == [] then + 0 + else + var finalIndex := |s| - 1; + SeqToNat(s[..finalIndex]) * 0x100 + s[finalIndex] as nat + } + + // By the following lemma, prepending a 0 to a byte sequence does not change its SeqToNat value + lemma SeqToNatZeroPrefix(s: seq) + ensures SeqToNat(s) == SeqToNat([0] + s) + { + if s == [] { + } else { + var s' := [0] + s; + var sLength := |s|; + var sFinalIndex := sLength - 1; + calc { + SeqToNat(s); + == + SeqToNat(s[..sFinalIndex]) * 0x100 + s[sFinalIndex] as nat; + == + SeqToNat([0] + s[..sFinalIndex]) * 0x100 + s[sFinalIndex] as nat; + == { assert (s'[..sLength] == [0] + s[..sFinalIndex]) && s'[sLength] == s[sFinalIndex]; } + SeqToNat(s'[..sLength]) * 0x100 + s'[sLength] as nat; + == + SeqToNat(s'); + == + SeqToNat([0] + s); + } + } + } + + // By the following lemma, SeqToNat(s) == n is automatically true if the given preconditions are true + lemma SeqWithUInt32Suffix(s: seq, n: nat) + requires n < UINT32_LIMIT + requires 4 <= |s| + requires var suffixStartIndex := |s| - 4; + (s[suffixStartIndex..] == UInt32ToSeq(n as uint32)) && + (forall i :: 0 <= i < suffixStartIndex ==> s[i] == 0) + ensures SeqToNat(s) == n + { + if |s| == 4 { + calc { + SeqToNat(s); + == + SeqToNat(s[..3]) + * 0x100 + s[3] as nat; + == { assert s[..3][..2] == s[..2] && s[..3][2] == s[2]; } + (SeqToNat(s[..2]) + * 0x100 + s[2] as nat) + * 0x100 + s[3] as nat; + == { assert s[..2][..1] == s[..1] && s[..2][1] == s[1]; } + ((SeqToNat(s[..1]) + * 0x100 + s[1] as nat) + * 0x100 + s[2] as nat) + * 0x100 + s[3] as nat; + == { assert s[..1][..0] == s[..0] && s[..1][0] == s[0]; } + (((SeqToNat(s[..0]) + * 0x100 + s[0] as nat) + * 0x100 + s[1] as nat) + * 0x100 + s[2] as nat) + * 0x100 + s[3] as nat; + == + n; + } + } else { + assert s == [0] + s[1..]; + SeqToNatZeroPrefix(s[1..]); + SeqWithUInt32Suffix(s[1..], n); + } + } +} diff --git a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy new file mode 100644 index 000000000..33a58eaa5 --- /dev/null +++ b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy @@ -0,0 +1,129 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +include "../../StandardLibrary/StandardLibrary.dfy" + include "../../Util/UTF8.dfy" + module {:extern "Dafny.Aws.Cryptography.StructuredEncryption.Types" } AwsCryptographyStructuredEncryptionTypes + { + import opened Wrappers + import opened StandardLibrary.UInt + import opened UTF8 + // Generic helpers for verification of mock/unit tests. + datatype DafnyCallEvent = DafnyCallEvent(input: I, output: O) + function Last(s: seq): T requires |s| > 0 { s[|s|-1] } + + // Begin Generated Types + + datatype CryptoAction = + | ENCRYPT_AND_SIGN + | SIGN_ONLY + | IGNORE + datatype CryptoSchema = | CryptoSchema ( + nameonly content: CryptoSchemaContent , + nameonly attributes: Option + ) + type CryptoSchemaAttributes = map + datatype CryptoSchemaContent = + | action(CryptoAction: CryptoAction) + | mapSchema(CryptoSchemaMap: CryptoSchemaMap) + | listSchema(CryptoSchemaList: CryptoSchemaList) + type CryptoSchemaList = seq + type CryptoSchemaMap = map + datatype DecryptStructureInput = | DecryptStructureInput ( + nameonly ciphertextStructure: StructuredData , + nameonly cryptoSchema: CryptoSchema + ) + datatype DecryptStructureOutput = | DecryptStructureOutput ( + nameonly plaintextStructure: StructuredData + ) + datatype EncryptStructureInput = | EncryptStructureInput ( + nameonly plaintextStructure: StructuredData , + nameonly cryptoSchema: CryptoSchema + ) + datatype EncryptStructureOutput = | EncryptStructureOutput ( + nameonly ciphertextStructure: StructuredData + ) + datatype StructuredData = | StructuredData ( + nameonly content: StructuredDataContent , + nameonly attributes: Option + ) + type StructuredDataAttributes = map + datatype StructuredDataContent = + | terminal(Terminal: Terminal) + | dataList(StructuredDataList: StructuredDataList) + | dataMap(StructuredDataMap: StructuredDataMap) + type StructuredDataList = seq + type StructuredDataMap = map + trait {:termination false} IStructuredEncryptionClient + { + ghost var EncryptStructureHistoricalCallEvents: seq>> + predicate EncryptStructureEnsuresPublicly(input: EncryptStructureInput, output: Result) + // The public method to be called by library consumers + method EncryptStructure ( input: EncryptStructureInput ) + returns (output: Result) + modifies this`EncryptStructureHistoricalCallEvents + ensures EncryptStructureEnsuresPublicly(input, output) + ensures + && 0 < |EncryptStructureHistoricalCallEvents| + && Last(EncryptStructureHistoricalCallEvents) == DafnyCallEvent(input, output) + + ghost var DecryptStructureHistoricalCallEvents: seq>> + predicate DecryptStructureEnsuresPublicly(input: DecryptStructureInput, output: Result) + // The public method to be called by library consumers + method DecryptStructure ( input: DecryptStructureInput ) + returns (output: Result) + modifies this`DecryptStructureHistoricalCallEvents + ensures DecryptStructureEnsuresPublicly(input, output) + ensures + && 0 < |DecryptStructureHistoricalCallEvents| + && Last(DecryptStructureHistoricalCallEvents) == DafnyCallEvent(input, output) + +} + datatype StructuredEncryptionConfig = | StructuredEncryptionConfig ( + + ) + type Terminal = seq + datatype Error = + // Local Error structures are listed here + + // Any dependent models are listed here + + // The Collection error is used to collect several errors together + // This is useful when composing OR logic. + // Consider the following method: + // + // method FN(n:I) + // returns (res: Result) + // ensures A(I).Success? ==> res.Success? + // ensures B(I).Success? ==> res.Success? + // ensures A(I).Failure? && B(I).Failure? ==> res.Failure? + // + // If either A || B is successful then FN is successful. + // And if A && B fail then FN will fail. + // But what information should FN transmit back to the caller? + // While it may be correct to hide these details from the caller, + // this can not be the globally correct option. + // Suppose that A and B can be blocked by different ACLs, + // and that their representation of I is only eventually consistent. + // How can the caller distinguish, at a minimum for logging, + // the difference between the four failure modes? + // || (!access(A(I)) && !access(B(I))) + // || (!exit(A(I)) && !exit(B(I))) + // || (!access(A(I)) && !exit(B(I))) + // || (!exit(A(I)) && !access(B(I))) + | Collection(list: seq) + // The Opaque error, used for native, extern, wrapped or unknown errors + | Opaque(obj: object) + type OpaqueError = e: Error | e.Opaque? witness * +} + abstract module AwsCryptographyStructuredEncryptionAbstract + { + import opened Wrappers + import opened StandardLibrary.UInt + import opened UTF8 + import opened Types = AwsCryptographyStructuredEncryptionTypes + function method DefaultStructuredEncryptionConfig(): StructuredEncryptionConfig + method StructuredEncryption(config: StructuredEncryptionConfig := DefaultStructuredEncryptionConfig()) + returns (res: Result) + ensures res.Success? ==> fresh(res.value) +} diff --git a/src/StructuredEncryption/model/StructuredEncryption.smithy b/src/StructuredEncryption/model/StructuredEncryption.smithy new file mode 100644 index 000000000..6d0e11149 --- /dev/null +++ b/src/StructuredEncryption/model/StructuredEncryption.smithy @@ -0,0 +1,136 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +namespace aws.cryptography.structuredEncryption + +use aws.polymorph#localService + +// TODO: Bikeshed on name +@aws.polymorph#localService( + sdkId: "StructuredEncryption", + config: StructuredEncryptionConfig, +) +service StructuredEncryption { + version: "2022-07-08", + operations: [EncryptStructure, DecryptStructure] +} + +structure StructuredEncryptionConfig { +} + +operation EncryptStructure { + input: EncryptStructureInput, + output: EncryptStructureOutput, +} + +operation DecryptStructure { + input: DecryptStructureInput, + output: DecryptStructureOutput, +} + +structure EncryptStructureInput { + @required + plaintextStructure: StructuredData, + @required + cryptoSchema: CryptoSchema + // CMM/Keyring + // EncryptionContext (stored and not stored) +} + +structure EncryptStructureOutput { + @required + ciphertextStructure: StructuredData +} + +structure DecryptStructureInput { + @required + ciphertextStructure: StructuredData, + @required + cryptoSchema: CryptoSchema + // CMM/Keyring + // EncryptionContext (stored and not stored) +} + +structure DecryptStructureOutput { + @required + plaintextStructure: StructuredData +} + +structure StructuredData { + // Each "node" in our structured data hold either + // a map of more data, a list of more data, or a terminal value + @required + content: StructuredDataContent, + // Each "node" in our structured data may additionally + // have a flat map to express something akin to XML attributes + attributes: StructuredDataAttributes +} + +union StructuredDataContent { + terminal: Terminal, + dataList: StructuredDataList, + dataMap: StructuredDataMap +} + +// This SDK only handles bytes. +// It is the reponsibility of the caller to +// serialize and deserialize the data they +// encrypt/decrypt with this SDK. +blob Terminal + +map StructuredDataMap { + key: String, + value: StructuredData +} + +list StructuredDataList { + member: StructuredData +} + +map StructuredDataAttributes { + key: String, + value: Terminal +} + +// This mimics the same structure as StructuredData above, +// only it's "leaves" are CryptoAction instead of Terminal. +structure CryptoSchema { + @required + content: CryptoSchemaContent, + attributes: CryptoSchemaAttributes +} + +union CryptoSchemaContent { + action: CryptoAction, + mapSchema: CryptoSchemaMap, + listSchema: CryptoSchemaList +} + +@enum([ + { + "name": "ENCRYPT_AND_SIGN", + "value": "ENCRYPT_AND_SIGN", + }, + { + "name": "SIGN_ONLY", + "value": "SIGN_ONLY", + }, + { + "name": "IGNORE", + "value": "IGNORE", + }, +]) +string CryptoAction + +map CryptoSchemaMap { + key: String, + value: CryptoSchema +} + +list CryptoSchemaList { + member: CryptoSchema +} + +map CryptoSchemaAttributes { + key: String, + value: CryptoAction +} \ No newline at end of file diff --git a/src/StructuredEncryption/model/traits.smithy b/src/StructuredEncryption/model/traits.smithy new file mode 100644 index 000000000..d7a176d74 --- /dev/null +++ b/src/StructuredEncryption/model/traits.smithy @@ -0,0 +1,8 @@ +namespace aws.polymorph + + +@trait(selector: "service") +structure localService { + sdkId: String, + config: String +} \ No newline at end of file diff --git a/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj b/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj new file mode 100644 index 000000000..af1cca8b7 --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj @@ -0,0 +1,80 @@ + + + + + netcoreapp3.1;net452 + 7.3 + false + false + + + + + + + + + + + runtime; build; native; contentfiles; analyzers; buildtransitive + all + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + \ No newline at end of file diff --git a/src/StructuredEncryption/src/StructuredEncryptionClient.dfy b/src/StructuredEncryption/src/StructuredEncryptionClient.dfy new file mode 100644 index 000000000..58e54838f --- /dev/null +++ b/src/StructuredEncryption/src/StructuredEncryptionClient.dfy @@ -0,0 +1,100 @@ +include "../../StandardLibrary/StandardLibrary.dfy" +include "../../StandardLibrary/UInt.dfy" +include "../model/AwsCryptographyStructuredEncryptionTypes.dfy" + +module {:extern "Dafny.Aws.StructuredEncryption.StructuredEncryptionClient"} StructuredEncryptionClient { + import opened Wrappers + import opened UInt = StandardLibrary.UInt + import opened StandardLibrary + import opened AwsCryptographyStructuredEncryptionTypes + + class StructuredEncryptionClient extends IStructuredEncryptionClient { + + constructor() {} + + predicate EncryptStructureEnsuresPublicly( + input: EncryptStructureInput, + output: Result) + { + true // TODO + } + + predicate DecryptStructureEnsuresPublicly( + input: DecryptStructureInput, + output: Result) + { + true // TODO + } + + method EncryptStructure(input: EncryptStructureInput) + returns (output: Result) + modifies this`EncryptStructureHistoricalCallEvents + ensures EncryptStructureEnsuresPublicly(input, output) + ensures + && 0 < |EncryptStructureHistoricalCallEvents| + && Last(EncryptStructureHistoricalCallEvents) == DafnyCallEvent(input, output) + { + // TODO: Currently stubbed out to return a hardcoded StructuredData + var stubbedBytes := [0x21, 0x64, 0x6c, 0x72, 0x6f, 0x77, 0x20, 0x2c, 0x6f, 0x6c, 0x6c, 0x65, 0x68]; + var stubbedStructure := StructuredData( + content := StructuredDataContent.dataMap( + StructuredDataMap := map[ + "foo" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "bar" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "fizzbuzz" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ) + ] + ), + attributes := None() + ); + + var encryptOutput := EncryptStructureOutput(ciphertextStructure := stubbedStructure); + + output := Success(encryptOutput); + EncryptStructureHistoricalCallEvents := EncryptStructureHistoricalCallEvents + [DafnyCallEvent(input, output)]; + } + + method DecryptStructure(input: DecryptStructureInput) + returns (output: Result) + modifies this`DecryptStructureHistoricalCallEvents + ensures DecryptStructureEnsuresPublicly(input, output) + ensures + && 0 < |DecryptStructureHistoricalCallEvents| + && Last(DecryptStructureHistoricalCallEvents) == DafnyCallEvent(input, output) + { + // TODO: Currently stubbed out to return a hardcoded StructuredData + var stubbedBytes := [0x68, 0x65, 0x6c, 0x6c, 0x6f, 0x2c, 0x20, 0x77, 0x6f, 0x72, 0x6c, 0x64]; + var stubbedStructure := StructuredData( + content := StructuredDataContent.dataMap( + StructuredDataMap := map[ + "foo" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "bar" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "fizzbuzz" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ) + ] + ), + attributes := None() + ); + + var decryptOutput := DecryptStructureOutput(plaintextStructure := stubbedStructure); + output := Success(decryptOutput); + DecryptStructureHistoricalCallEvents := DecryptStructureHistoricalCallEvents + [DafnyCallEvent(input, output)]; + } + } +} \ No newline at end of file diff --git a/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj b/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj new file mode 100644 index 000000000..a7240493d --- /dev/null +++ b/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj @@ -0,0 +1,30 @@ + + + + Library + net5.0 + false + + + + + + + + + 8 + + + + + + \ No newline at end of file diff --git a/src/StructuredEncryption/test/HappyCaseTests.dfy b/src/StructuredEncryption/test/HappyCaseTests.dfy new file mode 100644 index 000000000..47ab0d27c --- /dev/null +++ b/src/StructuredEncryption/test/HappyCaseTests.dfy @@ -0,0 +1,102 @@ +include "../../StandardLibrary/StandardLibrary.dfy" +include "../../StandardLibrary/UInt.dfy" +include "../src/StructuredEncryptionClient.dfy" +include "../model/AwsCryptographyStructuredEncryptionTypes.dfy" + +module HappyCaseTests { + import opened Wrappers + import opened UInt = StandardLibrary.UInt + import opened StandardLibrary + import opened AwsCryptographyStructuredEncryptionTypes + import StructuredEncryptionClient + + method {:test} TestEncryptStructure() { + var client := new StructuredEncryptionClient.StructuredEncryptionClient(); + + // This method is currently stubbed, so it doesn't matter what our input is + var inputStructure := StructuredData( + content := StructuredDataContent.terminal(Terminal := []), + attributes := None() + ); + var schema := CryptoSchema( + content := CryptoSchemaContent.action(CryptoAction := CryptoAction.ENCRYPT_AND_SIGN), + attributes := None() + ); + + var encryptRes := client.EncryptStructure( + EncryptStructureInput( + plaintextStructure:=inputStructure, + cryptoSchema:=schema + ) + ); + + var stubbedBytes := [0x21, 0x64, 0x6c, 0x72, 0x6f, 0x77, 0x20, 0x2c, 0x6f, 0x6c, 0x6c, 0x65, 0x68]; + var expectedValue := StructuredData( + content := StructuredDataContent.dataMap( + StructuredDataMap := map[ + "foo" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "bar" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "fizzbuzz" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ) + ] + ), + attributes := None() + ); + + expect encryptRes.Success?; + expect encryptRes.value.ciphertextStructure == expectedValue; + } + + method {:test} TestDecryptStructure() { + var client := new StructuredEncryptionClient.StructuredEncryptionClient(); + + // This method is currently stubbed, so it doesn't matter what our input is + var inputStructure := StructuredData( + content := StructuredDataContent.terminal(Terminal := []), + attributes := None() + ); + var schema := CryptoSchema( + content := CryptoSchemaContent.action(CryptoAction := CryptoAction.ENCRYPT_AND_SIGN), + attributes := None() + ); + + var encryptRes := client.DecryptStructure( + DecryptStructureInput( + ciphertextStructure:=inputStructure, + cryptoSchema:=schema + ) + ); + + var stubbedBytes := [0x68, 0x65, 0x6c, 0x6c, 0x6f, 0x2c, 0x20, 0x77, 0x6f, 0x72, 0x6c, 0x64]; + var expectedValue := StructuredData( + content := StructuredDataContent.dataMap( + StructuredDataMap := map[ + "foo" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "bar" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "fizzbuzz" := StructuredData( + content := StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ) + ] + ), + attributes := None() + ); + + expect encryptRes.Success?; + expect encryptRes.value.plaintextStructure == expectedValue; + } +} diff --git a/src/Util/UTF8.dfy b/src/Util/UTF8.dfy new file mode 100644 index 000000000..751d9bc63 --- /dev/null +++ b/src/Util/UTF8.dfy @@ -0,0 +1,154 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 + +include "../StandardLibrary/StandardLibrary.dfy" + +// Provides a function ValidUTF8 which checks if an array of bytes is a valid sequence of UTF8 code points. +// Each code point of a UTF8 string is one of the following variants, ranging from one to four bytes: +// Case 1 : 0xxxxxxx +// Case 2 : 110xxxxx 10xxxxxx +// Case 3 : 1110xxxx 10xxxxxx 10xxxxxx +// Case 4 : 11110xxx 10xxxxxx 10xxxxxx 10xxxxxx + +// The first uint8 of each code point is called the leading uint8, while the rest are called continuation bytes. + +// This does NOT perform any range checks on the values encoded. + +module {:extern "UTF8"} UTF8 { + import opened Wrappers + import opened UInt = StandardLibrary.UInt + + type ValidUTF8Bytes = i: seq | ValidUTF8Seq(i) witness [] + + // The tradeoff of assuming the external implementation of encode and decode is correct is worth the tradeoff + // of unlocking being able to express and hence prove so many other specifications + function method {:extern "Encode"} Encode(s: string): (res: Result) + // US-ASCII only needs a single UTF-8 byte per character + ensures IsASCIIString(s) ==> res.Success? && |res.value| == |s| + + function method {:extern "Decode"} Decode(b: ValidUTF8Bytes): (res: Result) + + // The implementation is exactly the same as the above, I just can't know if the input is valid or not though. + function method {:extern "AttemptDecode"} AttemptDecode(b: seq): (res: Result) + + predicate method IsASCIIString(s: string) { + forall i :: 0 <= i < |s| ==> s[i] as int < 128 + } + + predicate method Uses1Byte(s: seq) + requires |s| >= 1 + { + // Based on syntax detailed on https://tools.ietf.org/html/rfc3629#section-4 + 0x00 <= s[0] <= 0x7F + } + + predicate method Uses2Bytes(s: seq) + requires |s| >= 2 + { + // Based on syntax detailed on https://tools.ietf.org/html/rfc3629#section-4 + (0xC2 <= s[0] <= 0xDF) && (0x80 <= s[1] <= 0xBF) + } + + predicate method Uses3Bytes(s: seq) + requires |s| >= 3 + { + // Based on syntax detailed on https://tools.ietf.org/html/rfc3629#section-4 + ((s[0] == 0xE0) && (0xA0 <= s[1] <= 0xBF) && (0x80 <= s[2] <= 0xBF)) + || ((0xE1 <= s[0] <= 0xEC) && (0x80 <= s[1] <= 0xBF) && (0x80 <= s[2] <= 0xBF)) + || ((s[0] == 0xED) && (0x80 <= s[1] <= 0x9F) && (0x80 <= s[2] <= 0xBF)) + || ((0xEE <= s[0] <= 0xEF) && (0x80 <= s[1] <= 0xBF) && (0x80 <= s[2] <= 0xBF)) + } + + predicate method Uses4Bytes(s: seq) + requires |s| >= 4 + { + // Based on syntax detailed on https://tools.ietf.org/html/rfc3629#section-4 + ((s[0] == 0xF0) && (0x90 <= s[1] <= 0xBF) && (0x80 <= s[2] <= 0xBF) && (0x80 <= s[3] <= 0xBF)) + || ((0xF1 <= s[0] <= 0xF3) && (0x80 <= s[1] <= 0xBF) && (0x80 <= s[2] <= 0xBF) && (0x80 <= s[3] <= 0xBF)) + || ((s[0] == 0xF4) && (0x80 <= s[1] <= 0x8F) && (0x80 <= s[2] <= 0xBF) && (0x80 <= s[3] <= 0xBF)) + } + + predicate method {:tailrecursion} ValidUTF8Range(a: seq, lo: nat, hi: nat) + requires lo <= hi <= |a| + decreases hi - lo + { + if lo == hi then + true + else + var r := a[lo..hi]; + if Uses1Byte(r) then + ValidUTF8Range(a, lo + 1, hi) + else if 2 <= |r| && Uses2Bytes(r) then + ValidUTF8Range(a, lo + 2, hi) + else if 3 <= |r| && Uses3Bytes(r) then + ValidUTF8Range(a, lo + 3, hi) + else if 4 <= |r| && Uses4Bytes(r) then + ValidUTF8Range(a, lo + 4, hi) + else + false + } + + lemma ValidUTF8Embed(a: seq, b: seq, c: seq, lo: nat, hi: nat) + requires lo <= hi <= |b| + ensures ValidUTF8Range(b, lo, hi) == ValidUTF8Range(a + b + c, |a| + lo, |a| + hi) + decreases hi - lo + { + if lo == hi { + } else { + var r := b[lo..hi]; + var r' := (a + b + c)[|a| + lo..|a| + hi]; + assert r == r'; + if Uses1Byte(r) { + ValidUTF8Embed(a, b, c, lo + 1, hi); + } else if 2 <= |r| && Uses2Bytes(r) { + ValidUTF8Embed(a, b, c, lo + 2, hi); + } else if 3 <= |r| && Uses3Bytes(r) { + ValidUTF8Embed(a, b, c, lo + 3, hi); + } else if 4 <= |r| && Uses4Bytes(r) { + ValidUTF8Embed(a, b, c, lo + 4, hi); + } + } + } + + predicate method ValidUTF8Seq(s: seq) { + ValidUTF8Range(s, 0, |s|) + } + + lemma ValidUTF8Concat(s: seq, t: seq) + requires ValidUTF8Seq(s) && ValidUTF8Seq(t) + ensures ValidUTF8Seq(s + t) + { + var lo := 0; + while lo < |s| + invariant lo <= |s| + invariant ValidUTF8Range(s, lo, |s|) + invariant ValidUTF8Range(s + t, 0, |s + t|) == ValidUTF8Range(s + t, lo, |s + t|) + { + var r := (s + t)[lo..]; + if Uses1Byte(r) { + lo := lo + 1; + } else if 2 <= |r| && Uses2Bytes(r) { + lo := lo + 2; + } else if 3 <= |r| && Uses3Bytes(r) { + lo := lo + 3; + } else if 4 <= |r| && Uses4Bytes(r) { + lo := lo + 4; + } + } + calc { + ValidUTF8Seq(s + t); + == // def.ValidUTF8Seq + ValidUTF8Range(s + t, 0, |s + t|); + == // loop invariant + ValidUTF8Range(s + t, lo, |s + t|); + == { assert s + t == s + t + [] && lo == |s| && |s + t| == |s| + |t|; } + ValidUTF8Range(s + t + [], |s|, |s| + |t|); + == { ValidUTF8Embed(s, t, [], 0, |t|); } + ValidUTF8Range(t, 0, |t|); + == // def.ValidUTF8Seq + ValidUTF8Seq(t); + == // precondition + true; + } + } +} From e7e5c906b4450ec8e2e79ef03b74e1a9b548c3d0 Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Wed, 17 Aug 2022 16:18:59 -0700 Subject: [PATCH 02/13] Update Polymorph and add Makefile --- Makefile | 15 ++ README.md | 38 +-- polymorph | 2 +- ...sCryptographyStructuredEncryptionTypes.dfy | 66 +++++- .../model/StructuredEncryption.smithy | 5 +- .../runtimes/net/Generated/CryptoAction.cs | 19 ++ .../runtimes/net/Generated/CryptoSchema.cs | 28 +++ .../net/Generated/DecryptStructureInput.cs | 29 +++ .../net/Generated/DecryptStructureOutput.cs | 20 ++ .../net/Generated/EncryptStructureInput.cs | 29 +++ .../net/Generated/EncryptStructureOutput.cs | 20 ++ .../runtimes/net/Generated/OpaqueError.cs | 13 ++ .../runtimes/net/Generated/StructuredData.cs | 28 +++ .../net/Generated/StructuredEncryption.cs | 31 +++ .../Generated/StructuredEncryptionConfig.cs | 13 ++ .../runtimes/net/Generated/TypeConversion.cs | 221 ++++++++++++++++++ .../src/StructuredEncryptionClient.dfy | 38 +-- 17 files changed, 574 insertions(+), 41 deletions(-) create mode 100644 Makefile create mode 100644 src/StructuredEncryption/runtimes/net/Generated/CryptoAction.cs create mode 100644 src/StructuredEncryption/runtimes/net/Generated/CryptoSchema.cs create mode 100644 src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs create mode 100644 src/StructuredEncryption/runtimes/net/Generated/DecryptStructureOutput.cs create mode 100644 src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs create mode 100644 src/StructuredEncryption/runtimes/net/Generated/EncryptStructureOutput.cs create mode 100644 src/StructuredEncryption/runtimes/net/Generated/OpaqueError.cs create mode 100644 src/StructuredEncryption/runtimes/net/Generated/StructuredData.cs create mode 100644 src/StructuredEncryption/runtimes/net/Generated/StructuredEncryption.cs create mode 100644 src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionConfig.cs create mode 100644 src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs diff --git a/Makefile b/Makefile new file mode 100644 index 000000000..ef76ddc76 --- /dev/null +++ b/Makefile @@ -0,0 +1,15 @@ +build: + cd src/StructuredEncryption/src;\ + dotnet build + +test: build + cd src/StructuredEncryption/runtimes/net;\ + dotnet test -f netcoreapp3.1 Test + +generate-models: + cd polymorph/smithy-dotnet;\ + ./gradlew run --args="\ + -m src/StructuredEncryption/model\ + -namespace aws.cryptography.structuredEncryption\ + --output-dotnet src/StructuredEncryption/runtimes/net/Generated\ + -d src/StructuredEncryption/model" \ No newline at end of file diff --git a/README.md b/README.md index cb69d7eb6..3081f09fb 100644 --- a/README.md +++ b/README.md @@ -2,30 +2,42 @@ TODO: Edit your repository description on GitHub -### Rebuild the model +### Development Requirements + +Right now this repo has a bare-bones Makefile to help with some building/testing. +The Makefile doesn't currently build dev requirements, so you will need to grab these yourself. + +You will need at least: +- dotnet (I'm using v5.0.402) +- dafny (I'm using v3.4.0) -Currently only smithy -> dafny is working. Need to get smithy -> .NET with union support. +Additionally, I am working off a Mac and cannot guarantee that these commands will work on other platforms. -Run the following: +### Rebuild the model + +``` +make generate-models ``` -cd polymorph/smithy-dotnet -export MODEL_ROOT=/aws-dynamodb-encryption-dafny/src/StructuredEncryption/model -./gradlew run --args="\ - -m $MODEL_ROOT \ - -namespace aws.cryptography.structuredEncryption \ - -d $MODEL_ROOT" +Generates new models using the polymorph submodule. +Will overwrite models, but will not clean up old models that are not overwritten. + +### Build and Verify Dafny Code + +``` +make build ``` -### Run the tests +Builds and verifies the dafny code. -Assuming you are on a Mac: +### Run the tests ``` -cd src/StructuredEncryption/runtimes/net -dotnet test -f netcoreapp3.1 Test +make test ``` +Re-builds/re-verifies the dafny code and runs the dafny generated tests. + ## Security See [CONTRIBUTING](CONTRIBUTING.md#security-issue-notifications) for more information. diff --git a/polymorph b/polymorph index bff5c35aa..7787db256 160000 --- a/polymorph +++ b/polymorph @@ -1 +1 @@ -Subproject commit bff5c35aa3323ebe9b7ff38abcb27b01486d992e +Subproject commit 7787db256ee66bdf2a4e8fbdbca062a693b0cbe5 diff --git a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy index 33a58eaa5..c9b30b0e3 100644 --- a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy +++ b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy @@ -54,29 +54,69 @@ include "../../StandardLibrary/StandardLibrary.dfy" | dataMap(StructuredDataMap: StructuredDataMap) type StructuredDataList = seq type StructuredDataMap = map + class IStructuredEncryptionClientCallHistory { + ghost constructor() { + EncryptStructure := []; + DecryptStructure := []; +} + ghost var EncryptStructure: seq>> + ghost var DecryptStructure: seq>> +} trait {:termination false} IStructuredEncryptionClient { - ghost var EncryptStructureHistoricalCallEvents: seq>> + // Helper to define any additional modifies/reads clauses + // If your operations need to mutate state add it + // in your constructor function: + // Modifies := {your, fields, here, History}; + // If you do not need to mutate anything: + // Modifies := {History}; + ghost const Modifies: set + // For an unassigned const field defined in a trait, + // Dafny can only assign a value in the constructor. + // This means that for Dafny to reason about this value, + // it needs some way to know (an invariant), + // about the state of the object. + // This builds on the Valid/Repr paradigm + // To make this kind requires is safe to add + // to methods called from unverified code + // the predicate MUST NOT take any arguments. + // This means that the correctness of this requires + // MUST only be evaluated by the class itself. + // If you require any additional mutation, + // Then you MUST ensure everything you need in ValidState. + // You MUST also ensure ValidState in your constructor. + predicate ValidState() + ensures ValidState() ==> History in Modifies + ghost const History: IStructuredEncryptionClientCallHistory predicate EncryptStructureEnsuresPublicly(input: EncryptStructureInput, output: Result) // The public method to be called by library consumers method EncryptStructure ( input: EncryptStructureInput ) returns (output: Result) - modifies this`EncryptStructureHistoricalCallEvents - ensures EncryptStructureEnsuresPublicly(input, output) + requires + && ValidState() + modifies Modifies - {History} , + History`EncryptStructure + // Dafny will skip type parameters when generating a default decreases clause. + decreases Modifies ensures - && 0 < |EncryptStructureHistoricalCallEvents| - && Last(EncryptStructureHistoricalCallEvents) == DafnyCallEvent(input, output) + && ValidState() + ensures EncryptStructureEnsuresPublicly(input, output) + ensures History.EncryptStructure == old(History.EncryptStructure) + [DafnyCallEvent(input, output)] - ghost var DecryptStructureHistoricalCallEvents: seq>> predicate DecryptStructureEnsuresPublicly(input: DecryptStructureInput, output: Result) // The public method to be called by library consumers method DecryptStructure ( input: DecryptStructureInput ) returns (output: Result) - modifies this`DecryptStructureHistoricalCallEvents - ensures DecryptStructureEnsuresPublicly(input, output) + requires + && ValidState() + modifies Modifies - {History} , + History`DecryptStructure + // Dafny will skip type parameters when generating a default decreases clause. + decreases Modifies ensures - && 0 < |DecryptStructureHistoricalCallEvents| - && Last(DecryptStructureHistoricalCallEvents) == DafnyCallEvent(input, output) + && ValidState() + ensures DecryptStructureEnsuresPublicly(input, output) + ensures History.DecryptStructure == old(History.DecryptStructure) + [DafnyCallEvent(input, output)] } datatype StructuredEncryptionConfig = | StructuredEncryptionConfig ( @@ -125,5 +165,9 @@ include "../../StandardLibrary/StandardLibrary.dfy" function method DefaultStructuredEncryptionConfig(): StructuredEncryptionConfig method StructuredEncryption(config: StructuredEncryptionConfig := DefaultStructuredEncryptionConfig()) returns (res: Result) - ensures res.Success? ==> fresh(res.value) + ensures res.Success? ==> + && fresh(res.value) + && fresh(res.value.Modifies) + && fresh(res.value.History) + && res.value.ValidState() } diff --git a/src/StructuredEncryption/model/StructuredEncryption.smithy b/src/StructuredEncryption/model/StructuredEncryption.smithy index 6d0e11149..94fc151fc 100644 --- a/src/StructuredEncryption/model/StructuredEncryption.smithy +++ b/src/StructuredEncryption/model/StructuredEncryption.smithy @@ -4,7 +4,7 @@ namespace aws.cryptography.structuredEncryption use aws.polymorph#localService -// TODO: Bikeshed on name +// TODO: Bikeshed on name "StructuredEncryption" @aws.polymorph#localService( sdkId: "StructuredEncryption", config: StructuredEncryptionConfig, @@ -44,6 +44,7 @@ structure EncryptStructureOutput { structure DecryptStructureInput { @required ciphertextStructure: StructuredData, + // TODO the below should be a map @required cryptoSchema: CryptoSchema // CMM/Keyring @@ -71,7 +72,7 @@ union StructuredDataContent { dataMap: StructuredDataMap } -// This SDK only handles bytes. +// Only handles bytes. // It is the reponsibility of the caller to // serialize and deserialize the data they // encrypt/decrypt with this SDK. diff --git a/src/StructuredEncryption/runtimes/net/Generated/CryptoAction.cs b/src/StructuredEncryption/runtimes/net/Generated/CryptoAction.cs new file mode 100644 index 000000000..ab58f2a96 --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/CryptoAction.cs @@ -0,0 +1,19 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { + using Amazon.Runtime; public class CryptoAction : ConstantClass { + + + public static readonly CryptoAction ENCRYPT_AND_SIGN = new CryptoAction ("ENCRYPT_AND_SIGN"); + + public static readonly CryptoAction SIGN_ONLY = new CryptoAction ("SIGN_ONLY"); + + public static readonly CryptoAction IGNORE = new CryptoAction ("IGNORE"); + public static readonly CryptoAction [] Values = { + ENCRYPT_AND_SIGN , IGNORE , SIGN_ONLY +} ; + public CryptoAction (string value) : base(value) {} +} +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/CryptoSchema.cs b/src/StructuredEncryption/runtimes/net/Generated/CryptoSchema.cs new file mode 100644 index 000000000..ed7206998 --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/CryptoSchema.cs @@ -0,0 +1,28 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { + public class CryptoSchema { + private AWS.Cryptography.StructuredEncryption.CryptoSchemaContent _content ; + private System.Collections.Generic.Dictionary _attributes ; + public AWS.Cryptography.StructuredEncryption.CryptoSchemaContent Content { + get { return this._content; } + set { this._content = value; } +} + internal bool IsSetContent () { + return this._content != null; +} + public System.Collections.Generic.Dictionary Attributes { + get { return this._attributes; } + set { this._attributes = value; } +} + internal bool IsSetAttributes () { + return this._attributes != null; +} + public void Validate() { + if (!IsSetContent()) throw new System.ArgumentException("Missing value for required property 'Content'"); + +} +} +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs b/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs new file mode 100644 index 000000000..5be099e4b --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs @@ -0,0 +1,29 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { + public class DecryptStructureInput { + private AWS.Cryptography.StructuredEncryption.StructuredData _ciphertextStructure ; + private AWS.Cryptography.StructuredEncryption.CryptoSchema _cryptoSchema ; + public AWS.Cryptography.StructuredEncryption.StructuredData CiphertextStructure { + get { return this._ciphertextStructure; } + set { this._ciphertextStructure = value; } +} + internal bool IsSetCiphertextStructure () { + return this._ciphertextStructure != null; +} + public AWS.Cryptography.StructuredEncryption.CryptoSchema CryptoSchema { + get { return this._cryptoSchema; } + set { this._cryptoSchema = value; } +} + internal bool IsSetCryptoSchema () { + return this._cryptoSchema != null; +} + public void Validate() { + if (!IsSetCiphertextStructure()) throw new System.ArgumentException("Missing value for required property 'CiphertextStructure'"); + if (!IsSetCryptoSchema()) throw new System.ArgumentException("Missing value for required property 'CryptoSchema'"); + +} +} +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureOutput.cs b/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureOutput.cs new file mode 100644 index 000000000..32a637f94 --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureOutput.cs @@ -0,0 +1,20 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { + public class DecryptStructureOutput { + private AWS.Cryptography.StructuredEncryption.StructuredData _plaintextStructure ; + public AWS.Cryptography.StructuredEncryption.StructuredData PlaintextStructure { + get { return this._plaintextStructure; } + set { this._plaintextStructure = value; } +} + internal bool IsSetPlaintextStructure () { + return this._plaintextStructure != null; +} + public void Validate() { + if (!IsSetPlaintextStructure()) throw new System.ArgumentException("Missing value for required property 'PlaintextStructure'"); + +} +} +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs b/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs new file mode 100644 index 000000000..7f45b0f25 --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs @@ -0,0 +1,29 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { + public class EncryptStructureInput { + private AWS.Cryptography.StructuredEncryption.StructuredData _plaintextStructure ; + private AWS.Cryptography.StructuredEncryption.CryptoSchema _cryptoSchema ; + public AWS.Cryptography.StructuredEncryption.StructuredData PlaintextStructure { + get { return this._plaintextStructure; } + set { this._plaintextStructure = value; } +} + internal bool IsSetPlaintextStructure () { + return this._plaintextStructure != null; +} + public AWS.Cryptography.StructuredEncryption.CryptoSchema CryptoSchema { + get { return this._cryptoSchema; } + set { this._cryptoSchema = value; } +} + internal bool IsSetCryptoSchema () { + return this._cryptoSchema != null; +} + public void Validate() { + if (!IsSetPlaintextStructure()) throw new System.ArgumentException("Missing value for required property 'PlaintextStructure'"); + if (!IsSetCryptoSchema()) throw new System.ArgumentException("Missing value for required property 'CryptoSchema'"); + +} +} +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureOutput.cs b/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureOutput.cs new file mode 100644 index 000000000..a1b4af01f --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureOutput.cs @@ -0,0 +1,20 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { + public class EncryptStructureOutput { + private AWS.Cryptography.StructuredEncryption.StructuredData _ciphertextStructure ; + public AWS.Cryptography.StructuredEncryption.StructuredData CiphertextStructure { + get { return this._ciphertextStructure; } + set { this._ciphertextStructure = value; } +} + internal bool IsSetCiphertextStructure () { + return this._ciphertextStructure != null; +} + public void Validate() { + if (!IsSetCiphertextStructure()) throw new System.ArgumentException("Missing value for required property 'CiphertextStructure'"); + +} +} +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/OpaqueError.cs b/src/StructuredEncryption/runtimes/net/Generated/OpaqueError.cs new file mode 100644 index 000000000..be923ab7f --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/OpaqueError.cs @@ -0,0 +1,13 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { + public class OpaqueError : Exception { + public readonly object obj; + public OpaqueError(Exception ex) : base("OpaqueError:", ex) { this.obj = ex; } + public OpaqueError() : base("Unknown Unexpected Error") { } + public OpaqueError(object obj) : base("Opaque obj is not an Exception.") { this.obj = obj;} +} + +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/StructuredData.cs b/src/StructuredEncryption/runtimes/net/Generated/StructuredData.cs new file mode 100644 index 000000000..cc22c5b60 --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/StructuredData.cs @@ -0,0 +1,28 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { + public class StructuredData { + private AWS.Cryptography.StructuredEncryption.StructuredDataContent _content ; + private System.Collections.Generic.Dictionary _attributes ; + public AWS.Cryptography.StructuredEncryption.StructuredDataContent Content { + get { return this._content; } + set { this._content = value; } +} + internal bool IsSetContent () { + return this._content != null; +} + public System.Collections.Generic.Dictionary Attributes { + get { return this._attributes; } + set { this._attributes = value; } +} + internal bool IsSetAttributes () { + return this._attributes != null; +} + public void Validate() { + if (!IsSetContent()) throw new System.ArgumentException("Missing value for required property 'Content'"); + +} +} +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryption.cs b/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryption.cs new file mode 100644 index 000000000..d9f6a30ec --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryption.cs @@ -0,0 +1,31 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using System.IO; + using System.Collections.Generic; + using AWS.Cryptography.StructuredEncryption; + using Dafny.Aws.Cryptography.StructuredEncryption.Types; namespace AWS.Cryptography.StructuredEncryption { + public class StructuredEncryption { + private readonly IStructuredEncryptionClient _impl; + public StructuredEncryption(AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig input) + { + Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredEncryptionConfig internalInput = TypeConversion.ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_StructuredEncryptionConfig(input); + var result = Dafny.Aws.Cryptography.StructuredEncryption.__default.StructuredEncryption(internalInput); + if (result.is_Failure) throw TypeConversion.FromDafny_CommonError(result.dtor_error); + this._impl = result.dtor_value; +} + public AWS.Cryptography.StructuredEncryption.EncryptStructureOutput EncryptStructure(AWS.Cryptography.StructuredEncryption.EncryptStructureInput input) { + Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput internalInput = TypeConversion.ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput(input); + Wrappers_Compile._IResult result = _impl.EncryptStructure(internalInput); + if (result.is_Failure) throw TypeConversion.FromDafny_CommonError(result.dtor_error); + return TypeConversion.FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput(result.dtor_value); +} + public AWS.Cryptography.StructuredEncryption.DecryptStructureOutput DecryptStructure(AWS.Cryptography.StructuredEncryption.DecryptStructureInput input) { + Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureInput internalInput = TypeConversion.ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput(input); + Wrappers_Compile._IResult result = _impl.DecryptStructure(internalInput); + if (result.is_Failure) throw TypeConversion.FromDafny_CommonError(result.dtor_error); + return TypeConversion.FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput(result.dtor_value); +} +} +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionConfig.cs b/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionConfig.cs new file mode 100644 index 000000000..c2185994b --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionConfig.cs @@ -0,0 +1,13 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { + public class StructuredEncryptionConfig { + + + public void Validate() { + +} +} +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs b/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs new file mode 100644 index 000000000..e2ce05da2 --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs @@ -0,0 +1,221 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System.Linq; namespace AWS.Cryptography.StructuredEncryption { + internal static class TypeConversion { + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key (Dafny.ISequence value) { + return FromDafny_N6_smithy__N3_api__S6_String(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key (string value) { + return ToDafny_N6_smithy__N3_api__S6_String(value); +} + public static AWS.Cryptography.StructuredEncryption.DecryptStructureOutput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureOutput value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureOutput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureOutput)value; AWS.Cryptography.StructuredEncryption.DecryptStructureOutput converted = new AWS.Cryptography.StructuredEncryption.DecryptStructureOutput(); converted.PlaintextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure(concrete.plaintextStructure); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureOutput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput (AWS.Cryptography.StructuredEncryption.DecryptStructureOutput value) { + + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureOutput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure(value.PlaintextStructure) ) ; +} + public static AWS.Cryptography.StructuredEncryption.StructuredDataContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent value) { + +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent (AWS.Cryptography.StructuredEncryption.StructuredDataContent value) { + +} + public static AWS.Cryptography.StructuredEncryption.CryptoSchemaContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent(value); +} + public static string FromDafny_N6_smithy__N3_api__S6_String (Dafny.ISequence value) { + return new string(value.Elements); +} + public static Dafny.ISequence ToDafny_N6_smithy__N3_api__S6_String (string value) { + return Dafny.Sequence.FromString(value); +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes (Wrappers_Compile._IOption, Dafny.ISequence>> value) { + return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes(value.Extract()); +} + public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes (System.Collections.Generic.Dictionary value) { + return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes((System.Collections.Generic.Dictionary) value)); +} + public static System.IO.MemoryStream FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal (Dafny.ISequence value) { + return new System.IO.MemoryStream(value.Elements); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal (System.IO.MemoryStream value) { + return Dafny.Sequence.FromArray(value.ToArray()); +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes (Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction> value) { + return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value(pair.Cdr)); +} + public static Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes (System.Collections.Generic.Dictionary value) { + return Dafny.Map, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>.FromCollection(value.Select(pair => + new Dafny.Pair, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value(pair.Value)) +)); +} + public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData)value; AWS.Cryptography.StructuredEncryption.StructuredData converted = new AWS.Cryptography.StructuredEncryption.StructuredData(); converted.Content = (AWS.Cryptography.StructuredEncryption.StructuredDataContent) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content(concrete.content); + if (concrete.attributes.is_Some) converted.Attributes = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes(concrete.attributes); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData (AWS.Cryptography.StructuredEncryption.StructuredData value) { + System.Collections.Generic.Dictionary var_attributes = value.IsSetAttributes() ? value.Attributes : (System.Collections.Generic.Dictionary) null; + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content(value.Content) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes(var_attributes) ) ; +} + public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static AWS.Cryptography.StructuredEncryption.CryptoSchemaContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent value) { + +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent value) { + +} + public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static AWS.Cryptography.StructuredEncryption.CryptoAction FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value (AWS.Cryptography.StructuredEncryption.CryptoAction value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction(value); +} + public static AWS.Cryptography.StructuredEncryption.EncryptStructureOutput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureOutput value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureOutput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureOutput)value; AWS.Cryptography.StructuredEncryption.EncryptStructureOutput converted = new AWS.Cryptography.StructuredEncryption.EncryptStructureOutput(); converted.CiphertextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure(concrete.ciphertextStructure); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureOutput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput (AWS.Cryptography.StructuredEncryption.EncryptStructureOutput value) { + + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureOutput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure(value.CiphertextStructure) ) ; +} + public static AWS.Cryptography.StructuredEncryption.StructuredDataContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content (AWS.Cryptography.StructuredEncryption.StructuredDataContent value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent(value); +} + public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M12_cryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M12_cryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); +} + public static AWS.Cryptography.StructuredEncryption.CryptoAction FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction value) { + if (value.is_ENCRYPT__AND__SIGN) return AWS.Cryptography.StructuredEncryption.CryptoAction.ENCRYPT_AND_SIGN; + if (value.is_SIGN__ONLY) return AWS.Cryptography.StructuredEncryption.CryptoAction.SIGN_ONLY; + if (value.is_IGNORE) return AWS.Cryptography.StructuredEncryption.CryptoAction.IGNORE; +throw new System.ArgumentException("Invalid AWS.Cryptography.StructuredEncryption.CryptoAction value"); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction (AWS.Cryptography.StructuredEncryption.CryptoAction value) { + if (AWS.Cryptography.StructuredEncryption.CryptoAction.ENCRYPT_AND_SIGN.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_ENCRYPT__AND__SIGN(); + if (AWS.Cryptography.StructuredEncryption.CryptoAction.SIGN_ONLY.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_SIGN__ONLY(); + if (AWS.Cryptography.StructuredEncryption.CryptoAction.IGNORE.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_IGNORE(); +throw new System.ArgumentException("Invalid AWS.Cryptography.StructuredEncryption.CryptoAction value"); +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes (Dafny.IMap, Dafny.ISequence> value) { + return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value(pair.Cdr)); +} + public static Dafny.IMap, Dafny.ISequence> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes (System.Collections.Generic.Dictionary value) { + return Dafny.Map, Dafny.ISequence>.FromCollection(value.Select(pair => + new Dafny.Pair, Dafny.ISequence>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value(pair.Value)) +)); +} + public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); +} + public static AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_StructuredEncryptionConfig (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredEncryptionConfig value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig)value; AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig converted = new AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig(); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredEncryptionConfig ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_StructuredEncryptionConfig (AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig value) { + + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig ( ) ; +} + public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key (Dafny.ISequence value) { + return FromDafny_N6_smithy__N3_api__S6_String(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key (string value) { + return ToDafny_N6_smithy__N3_api__S6_String(value); +} + public static AWS.Cryptography.StructuredEncryption.EncryptStructureInput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput)value; AWS.Cryptography.StructuredEncryption.EncryptStructureInput converted = new AWS.Cryptography.StructuredEncryption.EncryptStructureInput(); converted.PlaintextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(concrete.plaintextStructure); + converted.CryptoSchema = (AWS.Cryptography.StructuredEncryption.CryptoSchema) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(concrete.cryptoSchema); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput (AWS.Cryptography.StructuredEncryption.EncryptStructureInput value) { + + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(value.PlaintextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(value.CryptoSchema) ) ; +} + public static System.IO.MemoryStream FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value (Dafny.ISequence value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value (System.IO.MemoryStream value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal(value); +} + public static AWS.Cryptography.StructuredEncryption.DecryptStructureInput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureInput value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput)value; AWS.Cryptography.StructuredEncryption.DecryptStructureInput converted = new AWS.Cryptography.StructuredEncryption.DecryptStructureInput(); converted.CiphertextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure(concrete.ciphertextStructure); + converted.CryptoSchema = (AWS.Cryptography.StructuredEncryption.CryptoSchema) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M12_cryptoSchema(concrete.cryptoSchema); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureInput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput (AWS.Cryptography.StructuredEncryption.DecryptStructureInput value) { + + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure(value.CiphertextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M12_cryptoSchema(value.CryptoSchema) ) ; +} + public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema)value; AWS.Cryptography.StructuredEncryption.CryptoSchema converted = new AWS.Cryptography.StructuredEncryption.CryptoSchema(); converted.Content = (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content(concrete.content); + if (concrete.attributes.is_Some) converted.Attributes = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes(concrete.attributes); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { + System.Collections.Generic.Dictionary var_attributes = value.IsSetAttributes() ? value.Attributes : (System.Collections.Generic.Dictionary) null; + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content(value.Content) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes(var_attributes) ) ; +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes (Wrappers_Compile._IOption, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>> value) { + return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes(value.Extract()); +} + public static Wrappers_Compile._IOption, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes (System.Collections.Generic.Dictionary value) { + return value == null ? Wrappers_Compile.Option, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>>.create_None() : Wrappers_Compile.Option, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes((System.Collections.Generic.Dictionary) value)); +} + public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static System.Exception FromDafny_CommonError(Dafny.Aws.Cryptography.StructuredEncryption.Types._IError value) { + switch(value) + { + + case Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Opaque dafnyVal: + return new OpaqueError(dafnyVal.obj); + default: + // The switch MUST be complete for _IError, so `value` MUST NOT be an _IError. (How did you get here?) + return new OpaqueError(); +} +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IError ToDafny_CommonError(System.Exception value) { + switch (value) + { + + // OpaqueError is redundant, but listed for completeness. + case OpaqueError exception: + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Opaque(exception); + case System.Exception exception: + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Opaque(exception); + default: + // The switch MUST be complete for System.Exception, so `value` MUST NOT be an System.Exception. (How did you get here?) + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Opaque(value); +} +} +} +} diff --git a/src/StructuredEncryption/src/StructuredEncryptionClient.dfy b/src/StructuredEncryption/src/StructuredEncryptionClient.dfy index 58e54838f..eabceec0a 100644 --- a/src/StructuredEncryption/src/StructuredEncryptionClient.dfy +++ b/src/StructuredEncryption/src/StructuredEncryptionClient.dfy @@ -9,30 +9,40 @@ module {:extern "Dafny.Aws.StructuredEncryption.StructuredEncryptionClient"} Str import opened AwsCryptographyStructuredEncryptionTypes class StructuredEncryptionClient extends IStructuredEncryptionClient { - - constructor() {} + predicate ValidState() + ensures ValidState() ==> History in Modifies + { + History in Modifies + } + + constructor() + ensures ValidState() && fresh(Modifies) && fresh(History) + { + History := new IStructuredEncryptionClientCallHistory(); + Modifies := {History}; + } predicate EncryptStructureEnsuresPublicly( input: EncryptStructureInput, output: Result) { - true // TODO + true } predicate DecryptStructureEnsuresPublicly( input: DecryptStructureInput, output: Result) { - true // TODO + true } method EncryptStructure(input: EncryptStructureInput) returns (output: Result) - modifies this`EncryptStructureHistoricalCallEvents + requires ValidState() + modifies Modifies - {History}, + History`EncryptStructure ensures EncryptStructureEnsuresPublicly(input, output) - ensures - && 0 < |EncryptStructureHistoricalCallEvents| - && Last(EncryptStructureHistoricalCallEvents) == DafnyCallEvent(input, output) + ensures History.EncryptStructure == old(History.EncryptStructure) + [DafnyCallEvent(input, output)] { // TODO: Currently stubbed out to return a hardcoded StructuredData var stubbedBytes := [0x21, 0x64, 0x6c, 0x72, 0x6f, 0x77, 0x20, 0x2c, 0x6f, 0x6c, 0x6c, 0x65, 0x68]; @@ -59,16 +69,16 @@ module {:extern "Dafny.Aws.StructuredEncryption.StructuredEncryptionClient"} Str var encryptOutput := EncryptStructureOutput(ciphertextStructure := stubbedStructure); output := Success(encryptOutput); - EncryptStructureHistoricalCallEvents := EncryptStructureHistoricalCallEvents + [DafnyCallEvent(input, output)]; + History.EncryptStructure := History.EncryptStructure + [DafnyCallEvent(input, output)]; } method DecryptStructure(input: DecryptStructureInput) returns (output: Result) - modifies this`DecryptStructureHistoricalCallEvents + requires ValidState() + modifies Modifies - {History}, + History`DecryptStructure ensures DecryptStructureEnsuresPublicly(input, output) - ensures - && 0 < |DecryptStructureHistoricalCallEvents| - && Last(DecryptStructureHistoricalCallEvents) == DafnyCallEvent(input, output) + ensures History.DecryptStructure == old(History.DecryptStructure) + [DafnyCallEvent(input, output)] { // TODO: Currently stubbed out to return a hardcoded StructuredData var stubbedBytes := [0x68, 0x65, 0x6c, 0x6c, 0x6f, 0x2c, 0x20, 0x77, 0x6f, 0x72, 0x6c, 0x64]; @@ -94,7 +104,7 @@ module {:extern "Dafny.Aws.StructuredEncryption.StructuredEncryptionClient"} Str var decryptOutput := DecryptStructureOutput(plaintextStructure := stubbedStructure); output := Success(decryptOutput); - DecryptStructureHistoricalCallEvents := DecryptStructureHistoricalCallEvents + [DafnyCallEvent(input, output)]; + History.DecryptStructure := History.DecryptStructure + [DafnyCallEvent(input, output)]; } } } \ No newline at end of file From d095caab011cf1e0dfd3962856fd9b0b59dcf41b Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Fri, 19 Aug 2022 19:03:46 -0700 Subject: [PATCH 03/13] Build with WIP MPL --- .gitmodules | 3 + Makefile | 14 +- README.md | 20 +- private-aws-encryption-sdk-dafny-staging | 1 + src/StandardLibrary/StandardLibrary.dfy | 406 ------------------ src/StandardLibrary/UInt.dfy | 332 -------------- ...sCryptographyStructuredEncryptionTypes.dfy | 60 ++- .../model/StructuredEncryption.smithy | 48 ++- src/StructuredEncryption/model/traits.smithy | 8 - .../net/Generated/DecryptStructureInput.cs | 46 +- .../net/Generated/EncryptStructureInput.cs | 32 ++ .../StructuredEncryptionException.cs | 9 + .../runtimes/net/Generated/TypeConversion.cs | 197 ++++++++- .../Test/AwsStructuredEncryptionTest.csproj | 63 +-- .../src/StructuredEncryptionClient.dfy | 94 ++-- .../src/StructuredEncryptionDafny.csproj | 38 ++ .../test/HappyCaseTests.dfy | 53 ++- src/Util/UTF8.dfy | 154 ------- 18 files changed, 559 insertions(+), 1019 deletions(-) create mode 160000 private-aws-encryption-sdk-dafny-staging delete mode 100644 src/StandardLibrary/StandardLibrary.dfy delete mode 100644 src/StandardLibrary/UInt.dfy delete mode 100644 src/StructuredEncryption/model/traits.smithy create mode 100644 src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionException.cs delete mode 100644 src/Util/UTF8.dfy diff --git a/.gitmodules b/.gitmodules index 27d8705bf..16c88e1f9 100644 --- a/.gitmodules +++ b/.gitmodules @@ -4,3 +4,6 @@ [submodule "polymorph"] path = polymorph url = git@github.com:awslabs/polymorph.git +[submodule "private-aws-encryption-sdk-dafny-staging"] + path = private-aws-encryption-sdk-dafny-staging + url = git@github.com:aws/private-aws-encryption-sdk-dafny-staging.git diff --git a/Makefile b/Makefile index ef76ddc76..13b416139 100644 --- a/Makefile +++ b/Makefile @@ -1,3 +1,6 @@ +ROOT_DIR:=$(shell dirname $(realpath $(firstword $(MAKEFILE_LIST)))) +DYLD_LIBRARY_PATH=$(DYLD_LIBRARY_PATH) + build: cd src/StructuredEncryption/src;\ dotnet build @@ -9,7 +12,10 @@ test: build generate-models: cd polymorph/smithy-dotnet;\ ./gradlew run --args="\ - -m src/StructuredEncryption/model\ - -namespace aws.cryptography.structuredEncryption\ - --output-dotnet src/StructuredEncryption/runtimes/net/Generated\ - -d src/StructuredEncryption/model" \ No newline at end of file + -m $(ROOT_DIR)/src/StructuredEncryption/Model \ + -namespace aws.cryptography.structuredEncryption \ + --output-dotnet $(ROOT_DIR)/src/StructuredEncryption/runtimes/net/Generated \ + -d $(ROOT_DIR)/private-aws-encryption-sdk-dafny-staging/src/AwsCryptographicMaterialProviders/Model \ + -d $(ROOT_DIR)/private-aws-encryption-sdk-dafny-staging/src/Crypto/Model \ + -d $(ROOT_DIR)/private-aws-encryption-sdk-dafny-staging/src/AWS-KMS/Model \ + -d $(ROOT_DIR)/private-aws-encryption-sdk-dafny-staging/model" \ No newline at end of file diff --git a/README.md b/README.md index 3081f09fb..ed67871f9 100644 --- a/README.md +++ b/README.md @@ -8,11 +8,27 @@ Right now this repo has a bare-bones Makefile to help with some building/testing The Makefile doesn't currently build dev requirements, so you will need to grab these yourself. You will need at least: -- dotnet (I'm using v5.0.402) -- dafny (I'm using v3.4.0) +- dotnet (I'm using v6.0.400) +- dafny (I'm using v3.7.3) Additionally, I am working off a Mac and cannot guarantee that these commands will work on other platforms. +#### System Dependencies - macOS only + +If you are using macOS then you must install OpenSSL 1.1, +and the OpenSSL 1.1 `lib` directory must be on the dynamic linker path at runtime. +We recommend that you install OpenSSL via Homebrew using `brew install openssl@1.1`, +and then set the `DYLD_LIBRARY_PATH` environment variable as follows: + +```bash +$ export DYLD_LIBRARY_PATH="/usr/local/opt/openssl@1.1/lib" +``` + +If the .NET runtime cannot locate your OpenSSL 1.1 libraries, +you may encounter an error that says: + +> No usable version of libssl was found + ### Rebuild the model ``` diff --git a/private-aws-encryption-sdk-dafny-staging b/private-aws-encryption-sdk-dafny-staging new file mode 160000 index 000000000..76353399b --- /dev/null +++ b/private-aws-encryption-sdk-dafny-staging @@ -0,0 +1 @@ +Subproject commit 76353399b1ab29d7b469e5df84844ff68ddec059 diff --git a/src/StandardLibrary/StandardLibrary.dfy b/src/StandardLibrary/StandardLibrary.dfy deleted file mode 100644 index a5dbf3abb..000000000 --- a/src/StandardLibrary/StandardLibrary.dfy +++ /dev/null @@ -1,406 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 - -include "../../libraries/src/Wrappers.dfy" -include "UInt.dfy" - -module StandardLibrary { - import opened Wrappers - import opened U = UInt - - function method {:tailrecursion} Join(ss: seq>, joiner: seq): (s: seq) - requires 0 < |ss| - { - if |ss| == 1 then ss[0] else ss[0] + joiner + Join(ss[1..], joiner) - } - - function method {:tailrecursion} Split(s: seq, delim: T): (res: seq>) - ensures delim !in s ==> res == [s] - ensures s == [] ==> res == [[]] - ensures 0 < |res| - ensures forall i :: 0 <= i < |res| ==> delim !in res[i] - ensures Join(res, [delim]) == s - decreases |s| - { - var i := FindIndexMatching(s, delim, 0); - if i.Some? then [s[..i.value]] + Split(s[(i.value + 1)..], delim) else [s] - } - - lemma WillSplitOnDelim(s: seq, delim: T, prefix: seq) - requires |prefix| < |s| - requires forall i :: 0 <= i < |prefix| ==> prefix[i] == s[i] - requires delim !in prefix && s[|prefix|] == delim - ensures Split(s, delim) == [prefix] + Split(s[|prefix| + 1..], delim) - { - calc { - Split(s, delim); - == - var i := FindIndexMatching(s, delim, 0); - if i.Some? then [s[..i.value]] + Split(s[i.value + 1..], delim) else [s]; - == { FindIndexMatchingLocatesElem(s, delim, 0, |prefix|); assert FindIndexMatching(s, delim, 0).Some?; } - [s[..|prefix|]] + Split(s[|prefix| + 1..], delim); - == { assert s[..|prefix|] == prefix; } - [prefix] + Split(s[|prefix| + 1..], delim); - } - } - - lemma WillNotSplitWithOutDelim(s: seq, delim: T) - requires delim !in s - ensures Split(s, delim) == [s] - { - calc { - Split(s, delim); - == - var i := FindIndexMatching(s, delim, 0); - if i.Some? then [s[..i.value]] + Split(s[i.value+1..], delim) else [s]; - == { FindIndexMatchingLocatesElem(s, delim, 0, |s|); } - [s]; - } - } - - lemma FindIndexMatchingLocatesElem(s: seq, c: T, start: nat, elemIndex: nat) - requires start <= elemIndex <= |s| - requires forall i :: start <= i < elemIndex ==> s[i] != c - requires elemIndex == |s| || s[elemIndex] == c - ensures FindIndexMatching(s, c, start) == if elemIndex == |s| then None else Some(elemIndex) - decreases elemIndex - start - {} - - function method FindIndexMatching(s: seq, c: T, i: nat): (index: Option) - requires i <= |s| - ensures index.Some? ==> i <= index.value < |s| && s[index.value] == c && c !in s[i..index.value] - ensures index.None? ==> c !in s[i..] - decreases |s| - i - { - FindIndex(s, x => x == c, i) - } - - function method {:tailrecursion} FindIndex(s: seq, f: T -> bool, i: nat): (index: Option) - requires i <= |s| - ensures index.Some? ==> i <= index.value < |s| && f(s[index.value]) && (forall j :: i <= j < index.value ==> !f(s[j])) - ensures index.None? ==> forall j :: i <= j < |s| ==> !f(s[j]) - decreases |s| - i - { - if i == |s| then None - else if f(s[i]) then Some(i) - else FindIndex(s, f, i + 1) - } - - function method {:tailrecursion} Filter(s: seq, f: T -> bool): (res: seq) - ensures forall i :: 0 <= i < |s| && f(s[i]) ==> s[i] in res - ensures forall i :: 0 <= i < |res| ==> res[i] in s && f(res[i]) - ensures |res| <= |s| - { - if |s| == 0 then [] - else if f(s[0]) then [s[0]] + Filter(s[1..], f) - else Filter(s[1..], f) - } - - lemma FilterIsDistributive(s: seq, s': seq, f: T -> bool) - ensures Filter(s + s', f) == Filter(s, f) + Filter(s', f) - { - if s == [] { - assert s + s' == s'; - } else { - var S := s + s'; - var s1 := s[1..]; - calc { - Filter(S, f); - == // def. Filter - if f(S[0]) then [S[0]] + Filter(S[1..], f) else Filter(S[1..], f); - == { assert S[0] == s[0] && S[1..] == s1 + s'; } - if f(s[0]) then [s[0]] + Filter(s1 + s', f) else Filter(s1 + s', f); - == { FilterIsDistributive(s1, s', f); } - if f(s[0]) then [s[0]] + (Filter(s1, f) + Filter(s', f)) else Filter(s1, f) + Filter(s', f); - == // associativity of + - if f(s[0]) then ([s[0]] + Filter(s1, f)) + Filter(s', f) else Filter(s1, f) + Filter(s', f); - == // distribute + over if-then-else - (if f(s[0]) then [s[0]] + Filter(s1, f) else Filter(s1, f)) + Filter(s', f); - == // def. Filter - Filter(s, f) + Filter(s', f); - } - } - } - - function method Min(a: int, b: int): int { - if a < b then a else b - } - - function method Fill(value: T, n: nat): seq - ensures |Fill(value, n)| == n - ensures forall i :: 0 <= i < n ==> Fill(value, n)[i] == value - { - seq(n, _ => value) - } - - method SeqToArray(s: seq) returns (a: array) - // "Fresh" expressions require editing memory - ensures fresh(a) - ensures a.Length == |s| - ensures forall i :: 0 <= i < |s| ==> a[i] == s[i] - { - a := new T[|s|](i requires 0 <= i < |s| => s[i]); - } - - lemma SeqPartsMakeWhole(s: seq, i: nat) - requires 0 <= i <= |s| - ensures s[..i] + s[i..] == s - {} - - /* - * Lexicographic comparison of sequences. - * - * Given two sequences `a` and `b` and a strict (that is, irreflexive) - * ordering `less` on the elements of these sequences, determine whether or not - * `a` is lexicographically "less than or equal to" `b`. - * - * `a` is lexicographically "less than or equal to" `b` holds iff - * there exists a `k` such that - * - the first `k` elements of `a` and `b` are the same - * - either: - * -- `a` has length `k` (that is, `a` is a prefix of `b`) - * -- `a[k]` is strictly less (using `less`) than `b[k]` - */ - - predicate method LexicographicLessOrEqual(a: seq, b: seq, less: (T, T) -> bool) { - exists k :: 0 <= k <= |a| && LexicographicLessOrEqualAux(a, b, less, k) - } - - predicate method LexicographicLessOrEqualAux(a: seq, b: seq, less: (T, T) -> bool, lengthOfCommonPrefix: nat) - requires 0 <= lengthOfCommonPrefix <= |a| - { - lengthOfCommonPrefix <= |b| - && (forall i :: 0 <= i < lengthOfCommonPrefix ==> a[i] == b[i]) - && (lengthOfCommonPrefix == |a| || (lengthOfCommonPrefix < |b| && less(a[lengthOfCommonPrefix], b[lengthOfCommonPrefix]))) - } - - /* - * In order for the lexicographic ordering above to have the expected properties, the - * relation "less" must be trichotomous and transitive. - * - * For an ordering `less` to be _trichotomous_ means that for any two `x` and `y`, - * EXACTLY one of the following three conditions holds: - * - less(x, y) - * - x == y - * - less(y, x) - * Note that being trichotomous implies being irreflexive. - */ - - predicate Trichotomous(less: (T, T) -> bool) { - (forall x, y :: less(x, y) || x == y || less(y, x)) && // at least one of the three - (forall x, y :: less(x, y) && less(y, x) ==> false) && // not both of the less's - (forall x, y :: less(x, y) ==> x != y) // not a less and the equality - } - - predicate Transitive(R: (T, T) -> bool) { - forall x, y, z :: R(x, y) && R(y, z) ==> R(x, z) - } - - /* - * Here is an example relation and a lemma that says the relation is appropriate for use in - * lexicographic orderings. - */ - - lemma UInt8LessIsTrichotomousTransitive() - ensures Trichotomous(UInt8Less) - ensures Transitive(UInt8Less) - { - } - - /* - * As the following lemmas show, the lexicographic ordering is reflexive, antisymmetric, transitive, and total. - * The proofs are a bit pedantic and include steps that can be automated. - */ - - lemma LexIsReflexive(a: seq, less: (T, T) -> bool) - ensures LexicographicLessOrEqual(a, a, less) - { - assert LexicographicLessOrEqualAux(a, a, less, |a|); - } - - lemma LexIsAntisymmetric(a: seq, b: seq, less: (T, T) -> bool) - requires Trich: Trichotomous(less) - requires LexicographicLessOrEqual(a, b, less) - requires LexicographicLessOrEqual(b, a, less) - ensures a == b - { - assert LessIrreflexive: forall x,y :: less(x, y) ==> x != y by { - reveal Trich; - } - assert ASymmetric: forall x,y :: less(x, y) && less(y, x) ==> false by { - reveal Trich; - } - var k0 :| 0 <= k0 <= |a| && LexicographicLessOrEqualAux(a, b, less, k0); - var k1 :| 0 <= k1 <= |b| && LexicographicLessOrEqualAux(b, a, less, k1); - var max := if k0 < k1 then k1 else k0; - assert max <= |a| && max <= |b|; - assert SameUntilMax: forall i :: 0 <= i < max ==> a[i] == b[i]; - assert AA: k0 == |a| || (k0 < |b| && less(a[k0], b[k0])); - assert BB: k1 == |b| || (k1 < |a| && less(b[k1], a[k1])); - calc { - true; - ==> { reveal AA, BB; } - (k0 == |a| || (k0 < |b| && less(a[k0], b[k0]))) && (k1 == |b| || (k1 < |a| && less(b[k1], a[k1]))); - == // distribute && and || - (k0 == |a| && k1 == |b|) || - (k0 == |a| && k1 < |a| && less(b[k1], a[k1])) || - (k0 < |b| && less(a[k0], b[k0]) && k1 == |b|) || - (k0 < |b| && less(a[k0], b[k0]) && k1 < |a| && less(b[k1], a[k1])); - == { reveal LessIrreflexive, SameUntilMax; } - (k0 == |a| && k1 == |b|) || - (k0 < |b| && less(a[k0], b[k0]) && k1 < |a| && less(b[k1], a[k1])); - ==> { reveal LessIrreflexive, SameUntilMax; assert max <= k0 && max <= k1; } - (k0 == |a| && k1 == |b|) || - (k0 < |b| && less(a[k0], b[k0]) && k1 < |a| && less(b[k1], a[k1]) && k0 == k1 == max); - == { reveal ASymmetric; } - k0 == |a| && k1 == |b|; - ==> { assert |a| == k0 <= max && |b| == k1 <= max ==> k0 == k1; } - max == |a| == |b|; - ==> { reveal SameUntilMax; } - a == b; - } - } - - lemma LexIsTransitive(a: seq, b: seq, c: seq, less: (T, T) -> bool) - requires Transitive(less) - requires LexicographicLessOrEqual(a, b, less) - requires LexicographicLessOrEqual(b, c, less) - ensures LexicographicLessOrEqual(a, c, less) - { - var k0 :| 0 <= k0 <= |a| && LexicographicLessOrEqualAux(a, b, less, k0); - var k1 :| 0 <= k1 <= |b| && LexicographicLessOrEqualAux(b, c, less, k1); - var k := if k0 < k1 then k0 else k1; - assert LexicographicLessOrEqualAux(a, c, less, k); - } - - lemma LexIsTotal(a: seq, b: seq, less: (T, T) -> bool) - requires Trich: Trichotomous(less) - ensures LexicographicLessOrEqual(a, b, less) || LexicographicLessOrEqual(b, a, less) - { - var m := 0; - while m < |a| && m < |b| && a[m] == b[m] - invariant m <= |a| && m <= |b| - invariant forall i :: 0 <= i < m ==> a[i] == b[i] - { - m := m + 1; - } - // m is the length of the common prefix of a and b - if m == |a| == |b| { - assert a == b; - LexIsReflexive(a, less); - } else if m == |a| < |b| { - assert LexicographicLessOrEqualAux(a, b, less, m); - } else if m == |b| < |a| { - assert LexicographicLessOrEqualAux(b, a, less, m); - } else { - assert m < |a| && m < |b|; - reveal Trich; - if - case less(a[m], b[m]) => - assert LexicographicLessOrEqualAux(a, b, less, m); - case less(b[m], a[m]) => - assert LexicographicLessOrEqualAux(b, a, less, m); - } - } - - /* - * SetToOrderedSequence(s, less) takes a set of T-strings and returns them as a sequence, - * ordered by the lexicographic ordering whose underlying irreflexive ordering is "less". - * The function is compilable, but will not exhibit enviable performance. - */ - - function method {:tailrecursion} SetToOrderedSequence(s: set>, less: (T, T) -> bool): (q: seq>) - requires Trichotomous(less) && Transitive(less) - ensures |s| == |q| - ensures forall i :: 0 <= i < |q| ==> q[i] in s - ensures forall k :: k in s ==> k in q - ensures forall i :: 0 < i < |q| ==> LexicographicLessOrEqual(q[i-1], q[i], less) - ensures forall i, j | 0 <= i < j < |q| :: q[i] != q[j]; - { - if s == {} then - [] - else - // In preparation for the assign-such-that statement below, we'll need to - // prove that a minimum exists and that it is unique. - // The following lemma shows existence: - ThereIsAMinimum(s, less); - // The following assertion shows uniqueness: - assert forall a, b :: IsMinimum(a, s, less) && IsMinimum(b, s, less) ==> a == b by { - // The proof of the assertion is the following forall statement. - // But why did we even bother to write the assert-by instead of - // just writing this forall statement in the first place? Because - // we are in an expression context and a forall statement cannot start - // an expression (because the "forall" makes the parser think that - // a forall expression is coming). - forall a, b | IsMinimum(a, s, less) && IsMinimum(b, s, less) { - // For the given a and b, the proof is settled by calling the following lemma: - MinimumIsUnique(a, b, s, less); - } - } - // The "a in s" in the following line follows from IsMinimum(a, s), so it - // is logically redundant. However, it is needed to convince the compiler - // that the assign-such-that statement is compilable. - var a :| a in s && IsMinimum(a, s, less); - [a] + SetToOrderedSequence(s - {a}, less) - } - - predicate method IsMinimum(a: seq, s: set>, less: (T, T) -> bool) { - a in s && - forall z :: z in s ==> LexicographicLessOrEqual(a, z, less) - } - - lemma ThereIsAMinimum(s: set>, less: (T, T) -> bool) - requires s != {} - requires Trichotomous(less) && Transitive(less) - ensures exists a :: IsMinimum(a, s, less) - { - var a := FindMinimum(s, less); - } - - lemma MinimumIsUnique(a: seq, b: seq, s: set>, less: (T, T) -> bool) - requires IsMinimum(a, s, less) && IsMinimum(b, s, less) - requires Trichotomous(less) - ensures a == b - { - LexIsAntisymmetric(a, b, less); - } - - lemma FindMinimum(s: set>, less: (T, T) -> bool) returns (a: seq) - requires s != {} - requires Trichotomous(less) && Transitive(less) - ensures IsMinimum(a, s, less) - { - a :| a in s; - if s == {a} { - LexIsReflexive(a, less); - } else { - var s' := s - {a}; - assert forall x :: x in s <==> x == a || x in s'; - var a' := FindMinimum(s', less); - if LexicographicLessOrEqual(a', a, less) { - a := a'; - } else { - assert LexicographicLessOrEqual(a, a', less) by { - LexIsTotal(a, a', less); - } - forall z | z in s - ensures LexicographicLessOrEqual(a, z, less) - { - if z == a { - LexIsReflexive(a, less); - } else { - calc { - true; - == // z in s && z != a - z in s'; - ==> // by postcondition of FindMinim(s') above - LexicographicLessOrEqual(a', z, less); - ==> { LexIsTransitive(a, a', z, less); } - LexicographicLessOrEqual(a, z, less); - } - } - } - } - } - } -} diff --git a/src/StandardLibrary/UInt.dfy b/src/StandardLibrary/UInt.dfy deleted file mode 100644 index 7f4613147..000000000 --- a/src/StandardLibrary/UInt.dfy +++ /dev/null @@ -1,332 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 - -module StandardLibrary.UInt { - - // TODO: Depend on types defined in dafny-lang/libraries instead - newtype uint8 = x | 0 <= x < 0x100 - newtype uint16 = x | 0 <= x < 0x1_0000 - newtype uint32 = x | 0 <= x < 0x1_0000_0000 - newtype uint64 = x | 0 <= x < 0x1_0000_0000_0000_0000 - - newtype int32 = x | -0x8000_0000 <= x < 0x8000_0000 - newtype int64 = x | -0x8000_0000_0000_0000 <= x < 0x8000_0000_0000_0000 - newtype posInt64 = x | 0 < x < 0x8000_0000_0000_0000 witness 1 - - const UINT16_LIMIT := 0x1_0000 - const UINT32_LIMIT := 0x1_0000_0000 - const UINT64_LIMIT := 0x1_0000_0000_0000_0000 - const INT32_MAX_LIMIT := 0x8000_0000 - const INT64_MAX_LIMIT := 0x8000_0000_0000_0000 - - predicate method UInt8Less(a: uint8, b: uint8) { a < b } - - predicate method HasUint16Len(s: seq) { - |s| < UINT16_LIMIT - } - - type seq16 = s: seq | HasUint16Len(s) - type Uint8Seq16 = seq16 - - predicate method HasUint32Len(s: seq) { - |s| < UINT32_LIMIT - } - - type seq32 = s: seq | HasUint32Len(s) - type Uint8Seq32 = seq32 - - predicate method HasUint64Len(s: seq) { - |s| < UINT64_LIMIT - } - - type seq64 = s: seq | HasUint64Len(s) - type Uint8Seq64 = seq64 - - function method UInt16ToSeq(x: uint16): (ret: seq) - ensures |ret| == 2 - ensures 0x100 * ret[0] as uint16 + ret[1] as uint16 == x - { - var b0 := (x / 0x100) as uint8; - var b1 := (x % 0x100) as uint8; - [b0, b1] - } - - function method SeqToUInt16(s: seq): (x: uint16) - requires |s| == 2 - ensures UInt16ToSeq(x) == s - ensures x >= 0 - { - var x0 := s[0] as uint16 * 0x100; - x0 + s[1] as uint16 - } - - lemma UInt16SeqSerializeDeserialize(x: uint16) - ensures SeqToUInt16(UInt16ToSeq(x)) == x - {} - - lemma UInt16SeqDeserializeSerialize(s: seq) - requires |s| == 2 - ensures UInt16ToSeq(SeqToUInt16(s)) == s - {} - - function method UInt32ToSeq(x: uint32): (ret: seq) - ensures |ret| == 4 - ensures 0x100_0000 * ret[0] as uint32 + 0x1_0000 * ret[1] as uint32 + 0x100 * ret[2] as uint32 + ret[3] as uint32 == x - { - var b0 := (x / 0x100_0000) as uint8; - var x0 := x - (b0 as uint32 * 0x100_0000); - - var b1 := (x0 / 0x1_0000) as uint8; - var x1 := x0 - (b1 as uint32 * 0x1_0000); - - var b2 := (x1 / 0x100) as uint8; - var b3 := (x1 % 0x100) as uint8; - [b0, b1, b2, b3] - } - - function method SeqToUInt32(s: seq): (x: uint32) - requires |s| == 4 - ensures UInt32ToSeq(x) == s - { - var x0 := s[0] as uint32 * 0x100_0000; - var x1 := x0 + s[1] as uint32 * 0x1_0000; - var x2 := x1 + s[2] as uint32 * 0x100; - x2 + s[3] as uint32 - } - - lemma UInt32SeqSerializeDeserialize(x: uint32) - ensures SeqToUInt32(UInt32ToSeq(x)) == x - {} - - lemma UInt32SeqDeserializeSerialize(s: seq) - requires |s| == 4 - ensures UInt32ToSeq(SeqToUInt32(s)) == s - {} - - function method UInt64ToSeq(x: uint64): (ret: seq) - ensures |ret| == 8 - ensures 0x100_0000_0000_0000 * ret[0] as uint64 + 0x1_0000_0000_0000 * ret[1] as uint64 + - 0x100_0000_0000 * ret[2] as uint64 + 0x1_0000_0000 * ret[3] as uint64 + 0x100_0000 * ret[4] as uint64 + - 0x1_0000 * ret[5] as uint64 + 0x100 * ret[6] as uint64 + ret[7] as uint64 == x - { - var b0 := (x / 0x100_0000_0000_0000) as uint8; - var x0 := x - (b0 as uint64 * 0x100_0000_0000_0000); - - var b1 := (x0 / 0x1_0000_0000_0000) as uint8; - var x1 := x0 - (b1 as uint64 * 0x1_0000_0000_0000); - - var b2 := (x1 / 0x100_0000_0000) as uint8; - var x2 := x1 - (b2 as uint64 * 0x100_0000_0000); - - var b3 := (x2 / 0x1_0000_0000) as uint8; - var x3 := x2 - (b3 as uint64 * 0x1_0000_0000); - - var b4 := (x3 / 0x100_0000) as uint8; - var x4 := x3 - (b4 as uint64 * 0x100_0000); - - var b5 := (x4 / 0x1_0000) as uint8; - var x5 := x4 - (b5 as uint64 * 0x1_0000); - - var b6 := (x5 / 0x100) as uint8; - var b7 := (x5 % 0x100) as uint8; - [b0, b1, b2, b3, b4, b5, b6, b7] - } - - function method SeqToUInt64(s: seq): (x: uint64) - requires |s| == 8 - ensures UInt64ToSeq(x) == s - { - var x0 := s[0] as uint64 * 0x100_0000_0000_0000; - var x1 := x0 + s[1] as uint64 * 0x1_0000_0000_0000; - var x2 := x1 + s[2] as uint64 * 0x100_0000_0000; - var x3 := x2 + s[3] as uint64 * 0x1_0000_0000; - var x4 := x3 + s[4] as uint64 * 0x100_0000; - var x5 := x4 + s[5] as uint64 * 0x1_0000; - var x6 := x5 + s[6] as uint64 * 0x100; - var x := x6 + s[7] as uint64; - UInt64SeqSerialize(x, s); - x - } - - lemma UInt64SeqSerialize(x: uint64, s: seq) - requires |s| == 8 - requires 0x100_0000_0000_0000 * s[0] as uint64 - + 0x1_0000_0000_0000 * s[1] as uint64 - + 0x100_0000_0000 * s[2] as uint64 - + 0x1_0000_0000 * s[3] as uint64 - + 0x100_0000 * s[4] as uint64 - + 0x1_0000 * s[5] as uint64 - + 0x100 * s[6] as uint64 - + s[7] as uint64 == x - ensures UInt64ToSeq(x) == s - { - calc { - UInt64ToSeq(x); - == - UInt64ToSeq(s[0] as uint64 * 0x100_0000_0000_0000 - + s[1] as uint64 * 0x1_0000_0000_0000 - + s[2] as uint64 * 0x100_0000_0000 - + s[3] as uint64 * 0x1_0000_0000 - + s[4] as uint64 * 0x100_0000 - + s[5] as uint64 * 0x1_0000 - + s[6] as uint64 * 0x100 - + s[7] as uint64); - == - var b0 := ((s[0] as uint64 * 0x100_0000_0000_0000 - + s[1] as uint64 * 0x1_0000_0000_0000 - + s[2] as uint64 * 0x100_0000_0000 - + s[3] as uint64 * 0x1_0000_0000 - + s[4] as uint64 * 0x100_0000 - + s[5] as uint64 * 0x1_0000 - + s[6] as uint64 * 0x100 - + s[7] as uint64) / 0x100_0000_0000_0000) as uint8; - assert b0 == s[0]; - var x0 := (s[0] as uint64 * 0x100_0000_0000_0000 - + s[1] as uint64 * 0x1_0000_0000_0000 - + s[2] as uint64 * 0x100_0000_0000 - + s[3] as uint64 * 0x1_0000_0000 - + s[4] as uint64 * 0x100_0000 - + s[5] as uint64 * 0x1_0000 - + s[6] as uint64 * 0x100 - + s[7] as uint64) - (b0 as uint64 * 0x100_0000_0000_0000); - assert x0 == (s[1] as uint64 * 0x1_0000_0000_0000 - + s[2] as uint64 * 0x100_0000_0000 - + s[3] as uint64 * 0x1_0000_0000 - + s[4] as uint64 * 0x100_0000 - + s[5] as uint64 * 0x1_0000 - + s[6] as uint64 * 0x100 - + s[7] as uint64); - - var b1 := (x0 / 0x1_0000_0000_0000) as uint8; - assert b1 == s[1]; - var x1 := x0 - (b1 as uint64 * 0x1_0000_0000_0000); - assert x1 == (s[2] as uint64 * 0x100_0000_0000 - + s[3] as uint64 * 0x1_0000_0000 - + s[4] as uint64 * 0x100_0000 - + s[5] as uint64 * 0x1_0000 - + s[6] as uint64 * 0x100 - + s[7] as uint64); - - var b2 := (x1 / 0x100_0000_0000) as uint8; - assert b2 == s[2]; - var x2 := x1 - (b2 as uint64 * 0x100_0000_0000); - assert x2 == (s[3] as uint64 * 0x1_0000_0000 - + s[4] as uint64 * 0x100_0000 - + s[5] as uint64 * 0x1_0000 - + s[6] as uint64 * 0x100 - + s[7] as uint64); - - var b3 := (x2 / 0x1_0000_0000) as uint8; - assert b3 == s[3]; - var x3 := x2 - (b3 as uint64 * 0x1_0000_0000); - assert x3 == (s[4] as uint64 * 0x100_0000 - + s[5] as uint64 * 0x1_0000 - + s[6] as uint64 * 0x100 - + s[7] as uint64); - - var b4 := (x3 / 0x100_0000) as uint8; - assert b4 == s[4]; - var x4 := x3 - (b4 as uint64 * 0x100_0000); - assert x4 == (s[5] as uint64 * 0x1_0000 - + s[6] as uint64 * 0x100 - + s[7] as uint64); - - var b5 := (x4 / 0x1_0000) as uint8; - assert b5 == s[5]; - var x5 := x4 - (b5 as uint64 * 0x1_0000); - assert x5 == (s[6] as uint64 * 0x100 + s[7] as uint64); - - var b6 := (x5 / 0x100) as uint8; - assert b6 == s[6]; - var b7 := (x5 % 0x100) as uint8; - assert b7 == s[7]; - [b0, b1, b2, b3, b4, b5, b6, b7]; - == - [s[0], s[1], s[2], s[3], s[4], s[5], s[6], s[7]]; - == - s; - } - } - - lemma UInt64SeqSerializeDeserialize(x: uint64) - ensures SeqToUInt64(UInt64ToSeq(x)) == x - {} - - lemma UInt64SeqDeserializeSerialize(s: seq) - requires |s| == 8 - ensures UInt64ToSeq(SeqToUInt64(s)) == s - {} - - function SeqToNat(s: seq): nat { - if s == [] then - 0 - else - var finalIndex := |s| - 1; - SeqToNat(s[..finalIndex]) * 0x100 + s[finalIndex] as nat - } - - // By the following lemma, prepending a 0 to a byte sequence does not change its SeqToNat value - lemma SeqToNatZeroPrefix(s: seq) - ensures SeqToNat(s) == SeqToNat([0] + s) - { - if s == [] { - } else { - var s' := [0] + s; - var sLength := |s|; - var sFinalIndex := sLength - 1; - calc { - SeqToNat(s); - == - SeqToNat(s[..sFinalIndex]) * 0x100 + s[sFinalIndex] as nat; - == - SeqToNat([0] + s[..sFinalIndex]) * 0x100 + s[sFinalIndex] as nat; - == { assert (s'[..sLength] == [0] + s[..sFinalIndex]) && s'[sLength] == s[sFinalIndex]; } - SeqToNat(s'[..sLength]) * 0x100 + s'[sLength] as nat; - == - SeqToNat(s'); - == - SeqToNat([0] + s); - } - } - } - - // By the following lemma, SeqToNat(s) == n is automatically true if the given preconditions are true - lemma SeqWithUInt32Suffix(s: seq, n: nat) - requires n < UINT32_LIMIT - requires 4 <= |s| - requires var suffixStartIndex := |s| - 4; - (s[suffixStartIndex..] == UInt32ToSeq(n as uint32)) && - (forall i :: 0 <= i < suffixStartIndex ==> s[i] == 0) - ensures SeqToNat(s) == n - { - if |s| == 4 { - calc { - SeqToNat(s); - == - SeqToNat(s[..3]) - * 0x100 + s[3] as nat; - == { assert s[..3][..2] == s[..2] && s[..3][2] == s[2]; } - (SeqToNat(s[..2]) - * 0x100 + s[2] as nat) - * 0x100 + s[3] as nat; - == { assert s[..2][..1] == s[..1] && s[..2][1] == s[1]; } - ((SeqToNat(s[..1]) - * 0x100 + s[1] as nat) - * 0x100 + s[2] as nat) - * 0x100 + s[3] as nat; - == { assert s[..1][..0] == s[..0] && s[..1][0] == s[0]; } - (((SeqToNat(s[..0]) - * 0x100 + s[0] as nat) - * 0x100 + s[1] as nat) - * 0x100 + s[2] as nat) - * 0x100 + s[3] as nat; - == - n; - } - } else { - assert s == [0] + s[1..]; - SeqToNatZeroPrefix(s[1..]); - SeqWithUInt32Suffix(s[1..], n); - } - } -} diff --git a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy index c9b30b0e3..3c3b33c2b 100644 --- a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy +++ b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy @@ -1,13 +1,17 @@ // Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. // SPDX-License-Identifier: Apache-2.0 // Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -include "../../StandardLibrary/StandardLibrary.dfy" - include "../../Util/UTF8.dfy" +// TODO had to manually update the below file. Polymorph needs to be fixed to not assuer location of this file +include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" +// TODO had to manually update the below file. Polymorph needs to be fixed to not assuer location of this file + include "../../../private-aws-encryption-sdk-dafny-staging/src/Util/UTF8.dfy" + include "../../../private-aws-encryption-sdk-dafny-staging/src/AwsCryptographicMaterialProviders/src/Index.dfy" module {:extern "Dafny.Aws.Cryptography.StructuredEncryption.Types" } AwsCryptographyStructuredEncryptionTypes { import opened Wrappers import opened StandardLibrary.UInt import opened UTF8 + import AwsCryptographyMaterialProvidersTypes // Generic helpers for verification of mock/unit tests. datatype DafnyCallEvent = DafnyCallEvent(input: I, output: O) function Last(s: seq): T requires |s| > 0 { s[|s|-1] } @@ -29,16 +33,26 @@ include "../../StandardLibrary/StandardLibrary.dfy" | listSchema(CryptoSchemaList: CryptoSchemaList) type CryptoSchemaList = seq type CryptoSchemaMap = map + type CryptoSchemas = map + type CryptoSchemaVersion = string datatype DecryptStructureInput = | DecryptStructureInput ( nameonly ciphertextStructure: StructuredData , - nameonly cryptoSchema: CryptoSchema + nameonly cryptoSchemas: CryptoSchemas , + nameonly keyring: Option , + nameonly cmm: Option , + nameonly implicitEncryptionContext: Option , + nameonly explicitEncryptionContext: Option ) datatype DecryptStructureOutput = | DecryptStructureOutput ( nameonly plaintextStructure: StructuredData ) datatype EncryptStructureInput = | EncryptStructureInput ( nameonly plaintextStructure: StructuredData , - nameonly cryptoSchema: CryptoSchema + nameonly cryptoSchema: CryptoSchema , + nameonly keyring: Option , + nameonly cmm: Option , + nameonly implicitEncryptionContext: Option , + nameonly explicitEncryptionContext: Option ) datatype EncryptStructureOutput = | EncryptStructureOutput ( nameonly ciphertextStructure: StructuredData @@ -93,11 +107,21 @@ include "../../StandardLibrary/StandardLibrary.dfy" method EncryptStructure ( input: EncryptStructureInput ) returns (output: Result) requires - && ValidState() + && ValidState() && ( input.keyring.Some? ==> + && input.keyring.value.ValidState() + && input.keyring.value.Modifies !! Modifies + ) && ( input.cmm.Some? ==> + && input.cmm.value.ValidState() + && input.cmm.value.Modifies !! Modifies + ) modifies Modifies - {History} , - History`EncryptStructure + History`EncryptStructure , + (if input.keyring.Some? then input.keyring.value.Modifies else {}) , + (if input.cmm.Some? then input.cmm.value.Modifies else {}) // Dafny will skip type parameters when generating a default decreases clause. - decreases Modifies + decreases Modifies , + (if input.keyring.Some? then input.keyring.value.Modifies else {}) , + (if input.cmm.Some? then input.cmm.value.Modifies else {}) ensures && ValidState() ensures EncryptStructureEnsuresPublicly(input, output) @@ -108,11 +132,21 @@ include "../../StandardLibrary/StandardLibrary.dfy" method DecryptStructure ( input: DecryptStructureInput ) returns (output: Result) requires - && ValidState() + && ValidState() && ( input.keyring.Some? ==> + && input.keyring.value.ValidState() + && input.keyring.value.Modifies !! Modifies + ) && ( input.cmm.Some? ==> + && input.cmm.value.ValidState() + && input.cmm.value.Modifies !! Modifies + ) modifies Modifies - {History} , - History`DecryptStructure + History`DecryptStructure , + (if input.keyring.Some? then input.keyring.value.Modifies else {}) , + (if input.cmm.Some? then input.cmm.value.Modifies else {}) // Dafny will skip type parameters when generating a default decreases clause. - decreases Modifies + decreases Modifies , + (if input.keyring.Some? then input.keyring.value.Modifies else {}) , + (if input.cmm.Some? then input.cmm.value.Modifies else {}) ensures && ValidState() ensures DecryptStructureEnsuresPublicly(input, output) @@ -125,9 +159,11 @@ include "../../StandardLibrary/StandardLibrary.dfy" type Terminal = seq datatype Error = // Local Error structures are listed here - + | StructuredEncryptionException ( + nameonly message: string + ) // Any dependent models are listed here - + | AwsCryptographyMaterialProviders(AwsCryptographyMaterialProviders: AwsCryptographyMaterialProvidersTypes.Error) // The Collection error is used to collect several errors together // This is useful when composing OR logic. // Consider the following method: diff --git a/src/StructuredEncryption/model/StructuredEncryption.smithy b/src/StructuredEncryption/model/StructuredEncryption.smithy index 94fc151fc..e1a48b481 100644 --- a/src/StructuredEncryption/model/StructuredEncryption.smithy +++ b/src/StructuredEncryption/model/StructuredEncryption.smithy @@ -2,6 +2,10 @@ // SPDX-License-Identifier: Apache-2.0 namespace aws.cryptography.structuredEncryption +use aws.cryptography.materialProviders#KeyringReference +use aws.cryptography.materialProviders#CryptographicMaterialsManagerReference +use aws.cryptography.materialProviders#EncryptionContext + use aws.polymorph#localService // TODO: Bikeshed on name "StructuredEncryption" @@ -11,7 +15,8 @@ use aws.polymorph#localService ) service StructuredEncryption { version: "2022-07-08", - operations: [EncryptStructure, DecryptStructure] + operations: [EncryptStructure, DecryptStructure], + errors: [StructuredEncryptionException] } structure StructuredEncryptionConfig { @@ -31,9 +36,14 @@ structure EncryptStructureInput { @required plaintextStructure: StructuredData, @required - cryptoSchema: CryptoSchema - // CMM/Keyring - // EncryptionContext (stored and not stored) + cryptoSchema: CryptoSchema, + + // A Keyring XOR a CMM MUST be specified + keyring: KeyringReference, + cmm: CryptographicMaterialsManagerReference, + + implicitEncryptionContext: EncryptionContext, + explicitEncryptionContext: EncryptionContext } structure EncryptStructureOutput { @@ -44,11 +54,15 @@ structure EncryptStructureOutput { structure DecryptStructureInput { @required ciphertextStructure: StructuredData, - // TODO the below should be a map @required - cryptoSchema: CryptoSchema - // CMM/Keyring - // EncryptionContext (stored and not stored) + cryptoSchemas: CryptoSchemas, + + // A Keyring XOR a CMM MUST be specified + keyring: KeyringReference, + cmm: CryptographicMaterialsManagerReference, + + implicitEncryptionContext: EncryptionContext, + explicitEncryptionContext: EncryptionContext } structure DecryptStructureOutput { @@ -57,7 +71,7 @@ structure DecryptStructureOutput { } structure StructuredData { - // Each "node" in our structured data hold either + // Each "node" in our structured data holds either // a map of more data, a list of more data, or a terminal value @required content: StructuredDataContent, @@ -134,4 +148,20 @@ list CryptoSchemaList { map CryptoSchemaAttributes { key: String, value: CryptoAction +} + +map CryptoSchemas { + key: CryptoSchemaVersion, + value: CryptoSchema +} + +string CryptoSchemaVersion + +///////////// +// Errors + +@error("client") +structure StructuredEncryptionException { + @required + message: String, } \ No newline at end of file diff --git a/src/StructuredEncryption/model/traits.smithy b/src/StructuredEncryption/model/traits.smithy deleted file mode 100644 index d7a176d74..000000000 --- a/src/StructuredEncryption/model/traits.smithy +++ /dev/null @@ -1,8 +0,0 @@ -namespace aws.polymorph - - -@trait(selector: "service") -structure localService { - sdkId: String, - config: String -} \ No newline at end of file diff --git a/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs b/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs index 5be099e4b..f53815088 100644 --- a/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs +++ b/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs @@ -5,7 +5,11 @@ using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { public class DecryptStructureInput { private AWS.Cryptography.StructuredEncryption.StructuredData _ciphertextStructure ; - private AWS.Cryptography.StructuredEncryption.CryptoSchema _cryptoSchema ; + private System.Collections.Generic.Dictionary _cryptoSchemas ; + private AWS.Cryptography.MaterialProviders.IKeyring _keyring ; + private AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager _cmm ; + private System.Collections.Generic.Dictionary _implicitEncryptionContext ; + private System.Collections.Generic.Dictionary _explicitEncryptionContext ; public AWS.Cryptography.StructuredEncryption.StructuredData CiphertextStructure { get { return this._ciphertextStructure; } set { this._ciphertextStructure = value; } @@ -13,16 +17,44 @@ public AWS.Cryptography.StructuredEncryption.StructuredData CiphertextStructure internal bool IsSetCiphertextStructure () { return this._ciphertextStructure != null; } - public AWS.Cryptography.StructuredEncryption.CryptoSchema CryptoSchema { - get { return this._cryptoSchema; } - set { this._cryptoSchema = value; } + public System.Collections.Generic.Dictionary CryptoSchemas { + get { return this._cryptoSchemas; } + set { this._cryptoSchemas = value; } } - internal bool IsSetCryptoSchema () { - return this._cryptoSchema != null; + internal bool IsSetCryptoSchemas () { + return this._cryptoSchemas != null; +} + public AWS.Cryptography.MaterialProviders.IKeyring Keyring { + get { return this._keyring; } + set { this._keyring = value; } +} + internal bool IsSetKeyring () { + return this._keyring != null; +} + public AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager Cmm { + get { return this._cmm; } + set { this._cmm = value; } +} + internal bool IsSetCmm () { + return this._cmm != null; +} + public System.Collections.Generic.Dictionary ImplicitEncryptionContext { + get { return this._implicitEncryptionContext; } + set { this._implicitEncryptionContext = value; } +} + internal bool IsSetImplicitEncryptionContext () { + return this._implicitEncryptionContext != null; +} + public System.Collections.Generic.Dictionary ExplicitEncryptionContext { + get { return this._explicitEncryptionContext; } + set { this._explicitEncryptionContext = value; } +} + internal bool IsSetExplicitEncryptionContext () { + return this._explicitEncryptionContext != null; } public void Validate() { if (!IsSetCiphertextStructure()) throw new System.ArgumentException("Missing value for required property 'CiphertextStructure'"); - if (!IsSetCryptoSchema()) throw new System.ArgumentException("Missing value for required property 'CryptoSchema'"); + if (!IsSetCryptoSchemas()) throw new System.ArgumentException("Missing value for required property 'CryptoSchemas'"); } } diff --git a/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs b/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs index 7f45b0f25..513421c97 100644 --- a/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs +++ b/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs @@ -6,6 +6,10 @@ public class EncryptStructureInput { private AWS.Cryptography.StructuredEncryption.StructuredData _plaintextStructure ; private AWS.Cryptography.StructuredEncryption.CryptoSchema _cryptoSchema ; + private AWS.Cryptography.MaterialProviders.IKeyring _keyring ; + private AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager _cmm ; + private System.Collections.Generic.Dictionary _implicitEncryptionContext ; + private System.Collections.Generic.Dictionary _explicitEncryptionContext ; public AWS.Cryptography.StructuredEncryption.StructuredData PlaintextStructure { get { return this._plaintextStructure; } set { this._plaintextStructure = value; } @@ -19,6 +23,34 @@ public AWS.Cryptography.StructuredEncryption.CryptoSchema CryptoSchema { } internal bool IsSetCryptoSchema () { return this._cryptoSchema != null; +} + public AWS.Cryptography.MaterialProviders.IKeyring Keyring { + get { return this._keyring; } + set { this._keyring = value; } +} + internal bool IsSetKeyring () { + return this._keyring != null; +} + public AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager Cmm { + get { return this._cmm; } + set { this._cmm = value; } +} + internal bool IsSetCmm () { + return this._cmm != null; +} + public System.Collections.Generic.Dictionary ImplicitEncryptionContext { + get { return this._implicitEncryptionContext; } + set { this._implicitEncryptionContext = value; } +} + internal bool IsSetImplicitEncryptionContext () { + return this._implicitEncryptionContext != null; +} + public System.Collections.Generic.Dictionary ExplicitEncryptionContext { + get { return this._explicitEncryptionContext; } + set { this._explicitEncryptionContext = value; } +} + internal bool IsSetExplicitEncryptionContext () { + return this._explicitEncryptionContext != null; } public void Validate() { if (!IsSetPlaintextStructure()) throw new System.ArgumentException("Missing value for required property 'PlaintextStructure'"); diff --git a/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionException.cs b/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionException.cs new file mode 100644 index 000000000..fa6fb4897 --- /dev/null +++ b/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionException.cs @@ -0,0 +1,9 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 +// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. +using System; + using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { + public class StructuredEncryptionException : Exception { + public StructuredEncryptionException(string message) : base(message) {} +} +} diff --git a/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs b/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs index e2ce05da2..8a34be042 100644 --- a/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs +++ b/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs @@ -3,11 +3,67 @@ // Do not modify this file. This file is machine generated, and any changes to it will be overwritten. using System.Linq; namespace AWS.Cryptography.StructuredEncryption { internal static class TypeConversion { + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_implicitEncryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { + return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); +} + public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_implicitEncryptionContext (System.Collections.Generic.Dictionary value) { + return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_explicitEncryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { + return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); +} + public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_explicitEncryptionContext (System.Collections.Generic.Dictionary value) { + return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); +} public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key (Dafny.ISequence value) { return FromDafny_N6_smithy__N3_api__S6_String(value); } public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key (string value) { return ToDafny_N6_smithy__N3_api__S6_String(value); +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext (Dafny.IMap, Dafny.ISequence> value) { + return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value(pair.Cdr)); +} + public static Dafny.IMap, Dafny.ISequence> ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext (System.Collections.Generic.Dictionary value) { + return Dafny.Map, Dafny.ISequence>.FromCollection(value.Select(pair => + new Dafny.Pair, Dafny.ISequence>(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value(pair.Value)) +)); +} + public static string FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes (Dafny.ISequence value) { + System.Text.UTF8Encoding utf8 = new System.Text.UTF8Encoding(false, true); +return utf8.GetString(value.Elements); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes (string value) { + System.Text.UTF8Encoding utf8 = new System.Text.UTF8Encoding(false, true); +return Dafny.Sequence.FromArray(utf8.GetBytes(value)); +} + public static AWS.Cryptography.StructuredEncryption.StructuredEncryptionException FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException (Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_StructuredEncryptionException value) { + return new AWS.Cryptography.StructuredEncryption.StructuredEncryptionException ( + FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message(value.message) + ) ; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_StructuredEncryptionException ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException (AWS.Cryptography.StructuredEncryption.StructuredEncryptionException value) { + + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_StructuredEncryptionException ( + ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message(value.Message) + ) ; +} + public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference (Dafny.Aws.Cryptography.MaterialProviders.Types.ICryptographicMaterialsManager value) { + if (value is NativeWrapper_CryptographicMaterialsManager nativeWrapper) return nativeWrapper._impl; +return new CryptographicMaterialsManager(value); + +} + public static Dafny.Aws.Cryptography.MaterialProviders.Types.ICryptographicMaterialsManager ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager value) { + switch (value) + { + case CryptographicMaterialsManager valueWithImpl: + return valueWithImpl._impl; + case CryptographicMaterialsManagerBase nativeImpl: + return new NativeWrapper_CryptographicMaterialsManager(nativeImpl); + default: + throw new System.ArgumentException( + "Custom implementations of CryptographicMaterialsManager must extend CryptographicMaterialsManagerBase."); +} } public static AWS.Cryptography.StructuredEncryption.DecryptStructureOutput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureOutput value) { Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureOutput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureOutput)value; AWS.Cryptography.StructuredEncryption.DecryptStructureOutput converted = new AWS.Cryptography.StructuredEncryption.DecryptStructureOutput(); converted.PlaintextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure(concrete.plaintextStructure); return converted; @@ -21,6 +77,12 @@ public static AWS.Cryptography.StructuredEncryption.StructuredDataContent FromDa } public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent (AWS.Cryptography.StructuredEncryption.StructuredDataContent value) { +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas (Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema> value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas(value); +} + public static Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas (System.Collections.Generic.Dictionary value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas(value); } public static AWS.Cryptography.StructuredEncryption.CryptoSchemaContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent value) { return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent(value); @@ -39,6 +101,12 @@ public static Dafny.ISequence ToDafny_N6_smithy__N3_api__S6_String (string } public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes (System.Collections.Generic.Dictionary value) { return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes((System.Collections.Generic.Dictionary) value)); +} + public static string FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value (Dafny.ISequence value) { + return FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value (string value) { + return ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); } public static System.IO.MemoryStream FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal (Dafny.ISequence value) { return new System.IO.MemoryStream(value.Elements); @@ -53,6 +121,12 @@ public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_struct return Dafny.Map, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>.FromCollection(value.Select(pair => new Dafny.Pair, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value(pair.Value)) )); +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_explicitEncryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { + return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); +} + public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_explicitEncryptionContext (System.Collections.Generic.Dictionary value) { + return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); } public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData)value; AWS.Cryptography.StructuredEncryption.StructuredData converted = new AWS.Cryptography.StructuredEncryption.StructuredData(); converted.Content = (AWS.Cryptography.StructuredEncryption.StructuredDataContent) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content(concrete.content); @@ -61,18 +135,47 @@ public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_ public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData (AWS.Cryptography.StructuredEncryption.StructuredData value) { System.Collections.Generic.Dictionary var_attributes = value.IsSetAttributes() ? value.Attributes : (System.Collections.Generic.Dictionary) null; return new Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content(value.Content) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes(var_attributes) ) ; +} + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key (Dafny.ISequence value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key (string value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion(value); } public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); } public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); } public static AWS.Cryptography.StructuredEncryption.CryptoSchemaContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent value) { } public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent value) { +} + public static AWS.Cryptography.MaterialProviders.IKeyring FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference (Dafny.Aws.Cryptography.MaterialProviders.Types.IKeyring value) { + if (value is NativeWrapper_Keyring nativeWrapper) return nativeWrapper._impl; +return new Keyring(value); + +} + public static Dafny.Aws.Cryptography.MaterialProviders.Types.IKeyring ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference (AWS.Cryptography.MaterialProviders.IKeyring value) { + switch (value) + { + case Keyring valueWithImpl: + return valueWithImpl._impl; + case KeyringBase nativeImpl: + return new NativeWrapper_Keyring(nativeImpl); + default: + throw new System.ArgumentException( + "Custom implementations of Keyring must extend KeyringBase."); +} } public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); @@ -92,18 +195,18 @@ public static AWS.Cryptography.StructuredEncryption.EncryptStructureOutput FromD public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureOutput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput (AWS.Cryptography.StructuredEncryption.EncryptStructureOutput value) { return new Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureOutput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure(value.CiphertextStructure) ) ; +} + public static string FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key (Dafny.ISequence value) { + return FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key (string value) { + return ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); } public static AWS.Cryptography.StructuredEncryption.StructuredDataContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent value) { return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent(value); } public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content (AWS.Cryptography.StructuredEncryption.StructuredDataContent value) { return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent(value); -} - public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M12_cryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M12_cryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); } public static AWS.Cryptography.StructuredEncryption.CryptoAction FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction value) { if (value.is_ENCRYPT__AND__SIGN) return AWS.Cryptography.StructuredEncryption.CryptoAction.ENCRYPT_AND_SIGN; @@ -149,14 +252,39 @@ public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryptio } public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key (string value) { return ToDafny_N6_smithy__N3_api__S6_String(value); +} + public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm (Wrappers_Compile._IOption value) { + return value.is_None ? (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference(value.Extract()); +} + public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager value) { + return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference((AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) value)); } public static AWS.Cryptography.StructuredEncryption.EncryptStructureInput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput value) { Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput)value; AWS.Cryptography.StructuredEncryption.EncryptStructureInput converted = new AWS.Cryptography.StructuredEncryption.EncryptStructureInput(); converted.PlaintextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(concrete.plaintextStructure); - converted.CryptoSchema = (AWS.Cryptography.StructuredEncryption.CryptoSchema) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(concrete.cryptoSchema); return converted; + converted.CryptoSchema = (AWS.Cryptography.StructuredEncryption.CryptoSchema) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(concrete.cryptoSchema); + if (concrete.keyring.is_Some) converted.Keyring = (AWS.Cryptography.MaterialProviders.IKeyring) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring(concrete.keyring); + if (concrete.cmm.is_Some) converted.Cmm = (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm(concrete.cmm); + if (concrete.implicitEncryptionContext.is_Some) converted.ImplicitEncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_implicitEncryptionContext(concrete.implicitEncryptionContext); + if (concrete.explicitEncryptionContext.is_Some) converted.ExplicitEncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_explicitEncryptionContext(concrete.explicitEncryptionContext); return converted; } public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput (AWS.Cryptography.StructuredEncryption.EncryptStructureInput value) { - - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(value.PlaintextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(value.CryptoSchema) ) ; + AWS.Cryptography.MaterialProviders.IKeyring var_keyring = value.IsSetKeyring() ? value.Keyring : (AWS.Cryptography.MaterialProviders.IKeyring) null; + AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager var_cmm = value.IsSetCmm() ? value.Cmm : (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null; + System.Collections.Generic.Dictionary var_implicitEncryptionContext = value.IsSetImplicitEncryptionContext() ? value.ImplicitEncryptionContext : (System.Collections.Generic.Dictionary) null; + System.Collections.Generic.Dictionary var_explicitEncryptionContext = value.IsSetExplicitEncryptionContext() ? value.ExplicitEncryptionContext : (System.Collections.Generic.Dictionary) null; + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(value.PlaintextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(value.CryptoSchema) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring(var_keyring) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm(var_cmm) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_implicitEncryptionContext(var_implicitEncryptionContext) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_explicitEncryptionContext(var_explicitEncryptionContext) ) ; +} + public static AWS.Cryptography.MaterialProviders.IKeyring FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring (Wrappers_Compile._IOption value) { + return value.is_None ? (AWS.Cryptography.MaterialProviders.IKeyring) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference(value.Extract()); +} + public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring (AWS.Cryptography.MaterialProviders.IKeyring value) { + return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference((AWS.Cryptography.MaterialProviders.IKeyring) value)); +} + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion (Dafny.ISequence value) { + return new string(value.Elements); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion (string value) { + return Dafny.Sequence.FromString(value); } public static System.IO.MemoryStream FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value (Dafny.ISequence value) { return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal(value); @@ -166,11 +294,42 @@ public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_struct } public static AWS.Cryptography.StructuredEncryption.DecryptStructureInput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureInput value) { Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput)value; AWS.Cryptography.StructuredEncryption.DecryptStructureInput converted = new AWS.Cryptography.StructuredEncryption.DecryptStructureInput(); converted.CiphertextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure(concrete.ciphertextStructure); - converted.CryptoSchema = (AWS.Cryptography.StructuredEncryption.CryptoSchema) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M12_cryptoSchema(concrete.cryptoSchema); return converted; + converted.CryptoSchemas = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas(concrete.cryptoSchemas); + if (concrete.keyring.is_Some) converted.Keyring = (AWS.Cryptography.MaterialProviders.IKeyring) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring(concrete.keyring); + if (concrete.cmm.is_Some) converted.Cmm = (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm(concrete.cmm); + if (concrete.implicitEncryptionContext.is_Some) converted.ImplicitEncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_implicitEncryptionContext(concrete.implicitEncryptionContext); + if (concrete.explicitEncryptionContext.is_Some) converted.ExplicitEncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_explicitEncryptionContext(concrete.explicitEncryptionContext); return converted; } public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureInput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput (AWS.Cryptography.StructuredEncryption.DecryptStructureInput value) { - - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure(value.CiphertextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M12_cryptoSchema(value.CryptoSchema) ) ; + AWS.Cryptography.MaterialProviders.IKeyring var_keyring = value.IsSetKeyring() ? value.Keyring : (AWS.Cryptography.MaterialProviders.IKeyring) null; + AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager var_cmm = value.IsSetCmm() ? value.Cmm : (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null; + System.Collections.Generic.Dictionary var_implicitEncryptionContext = value.IsSetImplicitEncryptionContext() ? value.ImplicitEncryptionContext : (System.Collections.Generic.Dictionary) null; + System.Collections.Generic.Dictionary var_explicitEncryptionContext = value.IsSetExplicitEncryptionContext() ? value.ExplicitEncryptionContext : (System.Collections.Generic.Dictionary) null; + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure(value.CiphertextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas(value.CryptoSchemas) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring(var_keyring) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm(var_cmm) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_implicitEncryptionContext(var_implicitEncryptionContext) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_explicitEncryptionContext(var_explicitEncryptionContext) ) ; +} + public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm (Wrappers_Compile._IOption value) { + return value.is_None ? (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference(value.Extract()); +} + public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager value) { + return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference((AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) value)); +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_implicitEncryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { + return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); +} + public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_implicitEncryptionContext (System.Collections.Generic.Dictionary value) { + return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); +} + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message (Dafny.ISequence value) { + return FromDafny_N6_smithy__N3_api__S6_String(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message (string value) { + return ToDafny_N6_smithy__N3_api__S6_String(value); +} + public static AWS.Cryptography.MaterialProviders.IKeyring FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring (Wrappers_Compile._IOption value) { + return value.is_None ? (AWS.Cryptography.MaterialProviders.IKeyring) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference(value.Extract()); +} + public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring (AWS.Cryptography.MaterialProviders.IKeyring value) { + return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference((AWS.Cryptography.MaterialProviders.IKeyring) value)); } public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema)value; AWS.Cryptography.StructuredEncryption.CryptoSchema converted = new AWS.Cryptography.StructuredEncryption.CryptoSchema(); converted.Content = (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content(concrete.content); @@ -179,6 +338,14 @@ public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aw public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { System.Collections.Generic.Dictionary var_attributes = value.IsSetAttributes() ? value.Attributes : (System.Collections.Generic.Dictionary) null; return new Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content(value.Content) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes(var_attributes) ) ; +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas (Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema> value) { + return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value(pair.Cdr)); +} + public static Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas (System.Collections.Generic.Dictionary value) { + return Dafny.Map, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema>.FromCollection(value.Select(pair => + new Dafny.Pair, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value(pair.Value)) +)); } public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes (Wrappers_Compile._IOption, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>> value) { return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes(value.Extract()); @@ -195,7 +362,8 @@ public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData public static System.Exception FromDafny_CommonError(Dafny.Aws.Cryptography.StructuredEncryption.Types._IError value) { switch(value) { - + case Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_StructuredEncryptionException dafnyVal: +return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException(dafnyVal); case Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Opaque dafnyVal: return new OpaqueError(dafnyVal.obj); default: @@ -206,7 +374,8 @@ public static System.Exception FromDafny_CommonError(Dafny.Aws.Cryptography.Stru public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IError ToDafny_CommonError(System.Exception value) { switch (value) { - + case AWS.Cryptography.StructuredEncryption.StructuredEncryptionException exception: + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException(exception); // OpaqueError is redundant, but listed for completeness. case OpaqueError exception: return new Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Opaque(exception); diff --git a/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj b/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj index af1cca8b7..998179110 100644 --- a/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj +++ b/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj @@ -7,18 +7,14 @@ See https://xunit.net/docs/why-no-netstandard. --> netcoreapp3.1;net452 - 7.3 + 10 false false - - - - - - - - - - - - - - - - - - + + + + + + @@ -65,6 +47,31 @@ + + + + + + + + + + + + + + + + + + + + + + + + + + + + @@ -27,4 +39,30 @@ --> + + + + + + + + + + + + + + + + + + + + + + + + + + \ No newline at end of file diff --git a/src/StructuredEncryption/test/HappyCaseTests.dfy b/src/StructuredEncryption/test/HappyCaseTests.dfy index 47ab0d27c..5713ea27b 100644 --- a/src/StructuredEncryption/test/HappyCaseTests.dfy +++ b/src/StructuredEncryption/test/HappyCaseTests.dfy @@ -1,17 +1,33 @@ -include "../../StandardLibrary/StandardLibrary.dfy" -include "../../StandardLibrary/UInt.dfy" +include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" include "../src/StructuredEncryptionClient.dfy" -include "../model/AwsCryptographyStructuredEncryptionTypes.dfy" +include "../Model/AwsCryptographyStructuredEncryptionTypes.dfy" +include "../../../private-aws-encryption-sdk-dafny-staging/src/AwsCryptographicMaterialProviders/src/Index.dfy" module HappyCaseTests { import opened Wrappers - import opened UInt = StandardLibrary.UInt import opened StandardLibrary import opened AwsCryptographyStructuredEncryptionTypes import StructuredEncryptionClient + import AwsCryptographyMaterialProvidersTypes + import MaterialProviders method {:test} TestEncryptStructure() { var client := new StructuredEncryptionClient.StructuredEncryptionClient(); + + // Create keyring. Currently doesn't matter what keyring we create. + var matProvRes := MaterialProviders.MaterialProviders(MaterialProviders.DefaultMaterialProvidersConfig()); + expect matProvRes.Success?; + var matProv := matProvRes.value; + + var keyringInput := AwsCryptographyMaterialProvidersTypes.CreateAwsKmsMrkMultiKeyringInput( + generator := Some("arn:aws:kms:us-west-2:658956600833:key/b3537ef1-d8dc-4780-9f5a-55776cbb2f7f"), + kmsKeyIds := None(), + clientSupplier := None(), + grantTokens := None() + ); + var keyringRes := matProv.CreateAwsKmsMrkMultiKeyring(keyringInput); + expect keyringRes.Success?; + var keyring := keyringRes.value; // This method is currently stubbed, so it doesn't matter what our input is var inputStructure := StructuredData( @@ -26,7 +42,11 @@ module HappyCaseTests { var encryptRes := client.EncryptStructure( EncryptStructureInput( plaintextStructure:=inputStructure, - cryptoSchema:=schema + cryptoSchema:=schema, + keyring:=Some(keyring), + cmm:=None(), + implicitEncryptionContext:=None(), + explicitEncryptionContext:=None() ) ); @@ -57,6 +77,21 @@ module HappyCaseTests { method {:test} TestDecryptStructure() { var client := new StructuredEncryptionClient.StructuredEncryptionClient(); + + // Create keyring. Currently doesn't matter what keyring we create. + var matProvRes := MaterialProviders.MaterialProviders(MaterialProviders.DefaultMaterialProvidersConfig()); + expect matProvRes.Success?; + var matProv := matProvRes.value; + + var keyringInput := AwsCryptographyMaterialProvidersTypes.CreateAwsKmsMrkMultiKeyringInput( + generator := Some("arn:aws:kms:us-west-2:658956600833:key/b3537ef1-d8dc-4780-9f5a-55776cbb2f7f"), + kmsKeyIds := None(), + clientSupplier := None(), + grantTokens := None() + ); + var keyringRes := matProv.CreateAwsKmsMrkMultiKeyring(keyringInput); + expect keyringRes.Success?; + var keyring := keyringRes.value; // This method is currently stubbed, so it doesn't matter what our input is var inputStructure := StructuredData( @@ -67,11 +102,17 @@ module HappyCaseTests { content := CryptoSchemaContent.action(CryptoAction := CryptoAction.ENCRYPT_AND_SIGN), attributes := None() ); + var schemaMap := map[]; + schemaMap := schemaMap["0":=schema]; var encryptRes := client.DecryptStructure( DecryptStructureInput( ciphertextStructure:=inputStructure, - cryptoSchema:=schema + cryptoSchemas:=schemaMap, + keyring:=Some(keyring), + cmm:=None(), + implicitEncryptionContext:=None(), + explicitEncryptionContext:=None() ) ); diff --git a/src/Util/UTF8.dfy b/src/Util/UTF8.dfy deleted file mode 100644 index 751d9bc63..000000000 --- a/src/Util/UTF8.dfy +++ /dev/null @@ -1,154 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 - -include "../StandardLibrary/StandardLibrary.dfy" - -// Provides a function ValidUTF8 which checks if an array of bytes is a valid sequence of UTF8 code points. -// Each code point of a UTF8 string is one of the following variants, ranging from one to four bytes: -// Case 1 : 0xxxxxxx -// Case 2 : 110xxxxx 10xxxxxx -// Case 3 : 1110xxxx 10xxxxxx 10xxxxxx -// Case 4 : 11110xxx 10xxxxxx 10xxxxxx 10xxxxxx - -// The first uint8 of each code point is called the leading uint8, while the rest are called continuation bytes. - -// This does NOT perform any range checks on the values encoded. - -module {:extern "UTF8"} UTF8 { - import opened Wrappers - import opened UInt = StandardLibrary.UInt - - type ValidUTF8Bytes = i: seq | ValidUTF8Seq(i) witness [] - - // The tradeoff of assuming the external implementation of encode and decode is correct is worth the tradeoff - // of unlocking being able to express and hence prove so many other specifications - function method {:extern "Encode"} Encode(s: string): (res: Result) - // US-ASCII only needs a single UTF-8 byte per character - ensures IsASCIIString(s) ==> res.Success? && |res.value| == |s| - - function method {:extern "Decode"} Decode(b: ValidUTF8Bytes): (res: Result) - - // The implementation is exactly the same as the above, I just can't know if the input is valid or not though. - function method {:extern "AttemptDecode"} AttemptDecode(b: seq): (res: Result) - - predicate method IsASCIIString(s: string) { - forall i :: 0 <= i < |s| ==> s[i] as int < 128 - } - - predicate method Uses1Byte(s: seq) - requires |s| >= 1 - { - // Based on syntax detailed on https://tools.ietf.org/html/rfc3629#section-4 - 0x00 <= s[0] <= 0x7F - } - - predicate method Uses2Bytes(s: seq) - requires |s| >= 2 - { - // Based on syntax detailed on https://tools.ietf.org/html/rfc3629#section-4 - (0xC2 <= s[0] <= 0xDF) && (0x80 <= s[1] <= 0xBF) - } - - predicate method Uses3Bytes(s: seq) - requires |s| >= 3 - { - // Based on syntax detailed on https://tools.ietf.org/html/rfc3629#section-4 - ((s[0] == 0xE0) && (0xA0 <= s[1] <= 0xBF) && (0x80 <= s[2] <= 0xBF)) - || ((0xE1 <= s[0] <= 0xEC) && (0x80 <= s[1] <= 0xBF) && (0x80 <= s[2] <= 0xBF)) - || ((s[0] == 0xED) && (0x80 <= s[1] <= 0x9F) && (0x80 <= s[2] <= 0xBF)) - || ((0xEE <= s[0] <= 0xEF) && (0x80 <= s[1] <= 0xBF) && (0x80 <= s[2] <= 0xBF)) - } - - predicate method Uses4Bytes(s: seq) - requires |s| >= 4 - { - // Based on syntax detailed on https://tools.ietf.org/html/rfc3629#section-4 - ((s[0] == 0xF0) && (0x90 <= s[1] <= 0xBF) && (0x80 <= s[2] <= 0xBF) && (0x80 <= s[3] <= 0xBF)) - || ((0xF1 <= s[0] <= 0xF3) && (0x80 <= s[1] <= 0xBF) && (0x80 <= s[2] <= 0xBF) && (0x80 <= s[3] <= 0xBF)) - || ((s[0] == 0xF4) && (0x80 <= s[1] <= 0x8F) && (0x80 <= s[2] <= 0xBF) && (0x80 <= s[3] <= 0xBF)) - } - - predicate method {:tailrecursion} ValidUTF8Range(a: seq, lo: nat, hi: nat) - requires lo <= hi <= |a| - decreases hi - lo - { - if lo == hi then - true - else - var r := a[lo..hi]; - if Uses1Byte(r) then - ValidUTF8Range(a, lo + 1, hi) - else if 2 <= |r| && Uses2Bytes(r) then - ValidUTF8Range(a, lo + 2, hi) - else if 3 <= |r| && Uses3Bytes(r) then - ValidUTF8Range(a, lo + 3, hi) - else if 4 <= |r| && Uses4Bytes(r) then - ValidUTF8Range(a, lo + 4, hi) - else - false - } - - lemma ValidUTF8Embed(a: seq, b: seq, c: seq, lo: nat, hi: nat) - requires lo <= hi <= |b| - ensures ValidUTF8Range(b, lo, hi) == ValidUTF8Range(a + b + c, |a| + lo, |a| + hi) - decreases hi - lo - { - if lo == hi { - } else { - var r := b[lo..hi]; - var r' := (a + b + c)[|a| + lo..|a| + hi]; - assert r == r'; - if Uses1Byte(r) { - ValidUTF8Embed(a, b, c, lo + 1, hi); - } else if 2 <= |r| && Uses2Bytes(r) { - ValidUTF8Embed(a, b, c, lo + 2, hi); - } else if 3 <= |r| && Uses3Bytes(r) { - ValidUTF8Embed(a, b, c, lo + 3, hi); - } else if 4 <= |r| && Uses4Bytes(r) { - ValidUTF8Embed(a, b, c, lo + 4, hi); - } - } - } - - predicate method ValidUTF8Seq(s: seq) { - ValidUTF8Range(s, 0, |s|) - } - - lemma ValidUTF8Concat(s: seq, t: seq) - requires ValidUTF8Seq(s) && ValidUTF8Seq(t) - ensures ValidUTF8Seq(s + t) - { - var lo := 0; - while lo < |s| - invariant lo <= |s| - invariant ValidUTF8Range(s, lo, |s|) - invariant ValidUTF8Range(s + t, 0, |s + t|) == ValidUTF8Range(s + t, lo, |s + t|) - { - var r := (s + t)[lo..]; - if Uses1Byte(r) { - lo := lo + 1; - } else if 2 <= |r| && Uses2Bytes(r) { - lo := lo + 2; - } else if 3 <= |r| && Uses3Bytes(r) { - lo := lo + 3; - } else if 4 <= |r| && Uses4Bytes(r) { - lo := lo + 4; - } - } - calc { - ValidUTF8Seq(s + t); - == // def.ValidUTF8Seq - ValidUTF8Range(s + t, 0, |s + t|); - == // loop invariant - ValidUTF8Range(s + t, lo, |s + t|); - == { assert s + t == s + t + [] && lo == |s| && |s + t| == |s| + |t|; } - ValidUTF8Range(s + t + [], |s|, |s| + |t|); - == { ValidUTF8Embed(s, t, [], 0, |t|); } - ValidUTF8Range(t, 0, |t|); - == // def.ValidUTF8Seq - ValidUTF8Seq(t); - == // precondition - true; - } - } -} From e4756438945e7d69cc4e2a7e23a884f70ba54a88 Mon Sep 17 00:00:00 2001 From: lavaleri <49660121+lavaleri@users.noreply.github.com> Date: Mon, 22 Aug 2022 09:52:14 -0700 Subject: [PATCH 04/13] Update src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy Co-authored-by: Tony Knapp <5892063+texastony@users.noreply.github.com> --- .../model/AwsCryptographyStructuredEncryptionTypes.dfy | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy index 3c3b33c2b..d59afe22f 100644 --- a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy +++ b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy @@ -1,7 +1,7 @@ // Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. // SPDX-License-Identifier: Apache-2.0 // Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -// TODO had to manually update the below file. Polymorph needs to be fixed to not assuer location of this file +// TODO had to manually update the below file. Polymorph needs to be fixed to not assume location of this file include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" // TODO had to manually update the below file. Polymorph needs to be fixed to not assuer location of this file include "../../../private-aws-encryption-sdk-dafny-staging/src/Util/UTF8.dfy" From 480e78b77adae65bf6fb469f6e60d8dd0ab3d0ed Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Mon, 22 Aug 2022 09:56:05 -0700 Subject: [PATCH 05/13] remove libraries submodule --- .gitmodules | 3 --- libraries | 1 - 2 files changed, 4 deletions(-) delete mode 160000 libraries diff --git a/.gitmodules b/.gitmodules index 16c88e1f9..3957403a2 100644 --- a/.gitmodules +++ b/.gitmodules @@ -1,6 +1,3 @@ -[submodule "libraries"] - path = libraries - url = git@github.com:dafny-lang/libraries.git [submodule "polymorph"] path = polymorph url = git@github.com:awslabs/polymorph.git diff --git a/libraries b/libraries deleted file mode 160000 index 735839fa6..000000000 --- a/libraries +++ /dev/null @@ -1 +0,0 @@ -Subproject commit 735839fa6e79ae60cc90fb529bf0cdb367e26b00 From c9dcdbeb3e4ca0e668e230cae8c4710ef25950ab Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Mon, 22 Aug 2022 09:56:49 -0700 Subject: [PATCH 06/13] Ugrade MPL submodule --- private-aws-encryption-sdk-dafny-staging | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/private-aws-encryption-sdk-dafny-staging b/private-aws-encryption-sdk-dafny-staging index 76353399b..7d69adaa0 160000 --- a/private-aws-encryption-sdk-dafny-staging +++ b/private-aws-encryption-sdk-dafny-staging @@ -1 +1 @@ -Subproject commit 76353399b1ab29d7b469e5df84844ff68ddec059 +Subproject commit 7d69adaa0ab3157d54965ff6cfad05993bc25080 From 38391cd808fe8a5ee85d1333f24b9688e8455a1c Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Mon, 22 Aug 2022 12:57:46 -0700 Subject: [PATCH 07/13] Refactor and get verified --- Makefile | 12 +- README.md | 14 +- ...sCryptographyStructuredEncryptionTypes.dfy | 9 +- .../model/StructuredEncryption.smithy | 17 +- .../net/Generated/DecryptStructureInput.cs | 20 +- .../net/Generated/EncryptStructureInput.cs | 24 +- .../runtimes/net/Generated/TypeConversion.cs | 362 +++++++++--------- src/StructuredEncryption/src/Index.dfy | 82 ++++ .../Operations/DecryptStructureOperation.dfy | 56 +++ .../Operations/EncryptStructureOperation.dfy | 61 +++ .../src/StructuredEncryptionClient.dfy | 130 ------- .../test/HappyCaseTests.dfy | 9 +- 12 files changed, 447 insertions(+), 349 deletions(-) create mode 100644 src/StructuredEncryption/src/Index.dfy create mode 100644 src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy create mode 100644 src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy delete mode 100644 src/StructuredEncryption/src/StructuredEncryptionClient.dfy diff --git a/Makefile b/Makefile index 13b416139..b254cb7bd 100644 --- a/Makefile +++ b/Makefile @@ -1,11 +1,19 @@ ROOT_DIR:=$(shell dirname $(realpath $(firstword $(MAKEFILE_LIST)))) DYLD_LIBRARY_PATH=$(DYLD_LIBRARY_PATH) -build: +verify: + cd src/StructuredEncryption/src;\ + dotnet build -t:VerifyDafny + +build: build-dotnet + +test: test-dotnet + +build-dotnet: cd src/StructuredEncryption/src;\ dotnet build -test: build +test-dotnet: build cd src/StructuredEncryption/runtimes/net;\ dotnet test -f netcoreapp3.1 Test diff --git a/README.md b/README.md index ed67871f9..ba36e5d97 100644 --- a/README.md +++ b/README.md @@ -38,13 +38,21 @@ make generate-models Generates new models using the polymorph submodule. Will overwrite models, but will not clean up old models that are not overwritten. -### Build and Verify Dafny Code +### Verify Dafny Code + +``` +make verify +``` + +Verifies the entire project. + +### Build Dafny Code ``` make build ``` -Builds and verifies the dafny code. +Compiles the dafny code into target languages (currently just .NET) and builds. ### Run the tests @@ -52,7 +60,7 @@ Builds and verifies the dafny code. make test ``` -Re-builds/re-verifies the dafny code and runs the dafny generated tests. +Runs the dafny generated tests. ## Security diff --git a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy index d59afe22f..606cd35c5 100644 --- a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy +++ b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy @@ -40,19 +40,19 @@ include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/S nameonly cryptoSchemas: CryptoSchemas , nameonly keyring: Option , nameonly cmm: Option , - nameonly implicitEncryptionContext: Option , - nameonly explicitEncryptionContext: Option + nameonly encryptionContext: Option ) datatype DecryptStructureOutput = | DecryptStructureOutput ( nameonly plaintextStructure: StructuredData ) + type EncryptionContextFieldList = seq datatype EncryptStructureInput = | EncryptStructureInput ( nameonly plaintextStructure: StructuredData , nameonly cryptoSchema: CryptoSchema , nameonly keyring: Option , nameonly cmm: Option , - nameonly implicitEncryptionContext: Option , - nameonly explicitEncryptionContext: Option + nameonly encryptionContext: Option , + nameonly requiredContextFieldsOnDecrypt: Option ) datatype EncryptStructureOutput = | EncryptStructureOutput ( nameonly ciphertextStructure: StructuredData @@ -157,6 +157,7 @@ include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/S ) type Terminal = seq + type Utf8Bytes = ValidUTF8Bytes datatype Error = // Local Error structures are listed here | StructuredEncryptionException ( diff --git a/src/StructuredEncryption/model/StructuredEncryption.smithy b/src/StructuredEncryption/model/StructuredEncryption.smithy index e1a48b481..ec1c739c4 100644 --- a/src/StructuredEncryption/model/StructuredEncryption.smithy +++ b/src/StructuredEncryption/model/StructuredEncryption.smithy @@ -7,6 +7,10 @@ use aws.cryptography.materialProviders#CryptographicMaterialsManagerReference use aws.cryptography.materialProviders#EncryptionContext use aws.polymorph#localService +use aws.polymorph#dafnyUtf8Bytes + +@dafnyUtf8Bytes +string Utf8Bytes // TODO: Bikeshed on name "StructuredEncryption" @aws.polymorph#localService( @@ -42,8 +46,8 @@ structure EncryptStructureInput { keyring: KeyringReference, cmm: CryptographicMaterialsManagerReference, - implicitEncryptionContext: EncryptionContext, - explicitEncryptionContext: EncryptionContext + encryptionContext: EncryptionContext, + contextFieldsRequiredOnDecrypt: EncryptionContextFieldList } structure EncryptStructureOutput { @@ -61,8 +65,7 @@ structure DecryptStructureInput { keyring: KeyringReference, cmm: CryptographicMaterialsManagerReference, - implicitEncryptionContext: EncryptionContext, - explicitEncryptionContext: EncryptionContext + encryptionContext: EncryptionContext, } structure DecryptStructureOutput { @@ -70,6 +73,12 @@ structure DecryptStructureOutput { plaintextStructure: StructuredData } +// TODO move to MPL +// TODO this is better represented as a set +list EncryptionContextFieldList { + member: Utf8Bytes +} + structure StructuredData { // Each "node" in our structured data holds either // a map of more data, a list of more data, or a terminal value diff --git a/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs b/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs index f53815088..ad1e7cfaf 100644 --- a/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs +++ b/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs @@ -8,8 +8,7 @@ public class DecryptStructureInput { private System.Collections.Generic.Dictionary _cryptoSchemas ; private AWS.Cryptography.MaterialProviders.IKeyring _keyring ; private AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager _cmm ; - private System.Collections.Generic.Dictionary _implicitEncryptionContext ; - private System.Collections.Generic.Dictionary _explicitEncryptionContext ; + private System.Collections.Generic.Dictionary _encryptionContext ; public AWS.Cryptography.StructuredEncryption.StructuredData CiphertextStructure { get { return this._ciphertextStructure; } set { this._ciphertextStructure = value; } @@ -38,19 +37,12 @@ public AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager Cmm { internal bool IsSetCmm () { return this._cmm != null; } - public System.Collections.Generic.Dictionary ImplicitEncryptionContext { - get { return this._implicitEncryptionContext; } - set { this._implicitEncryptionContext = value; } + public System.Collections.Generic.Dictionary EncryptionContext { + get { return this._encryptionContext; } + set { this._encryptionContext = value; } } - internal bool IsSetImplicitEncryptionContext () { - return this._implicitEncryptionContext != null; -} - public System.Collections.Generic.Dictionary ExplicitEncryptionContext { - get { return this._explicitEncryptionContext; } - set { this._explicitEncryptionContext = value; } -} - internal bool IsSetExplicitEncryptionContext () { - return this._explicitEncryptionContext != null; + internal bool IsSetEncryptionContext () { + return this._encryptionContext != null; } public void Validate() { if (!IsSetCiphertextStructure()) throw new System.ArgumentException("Missing value for required property 'CiphertextStructure'"); diff --git a/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs b/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs index 513421c97..a230d4799 100644 --- a/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs +++ b/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs @@ -8,8 +8,8 @@ public class EncryptStructureInput { private AWS.Cryptography.StructuredEncryption.CryptoSchema _cryptoSchema ; private AWS.Cryptography.MaterialProviders.IKeyring _keyring ; private AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager _cmm ; - private System.Collections.Generic.Dictionary _implicitEncryptionContext ; - private System.Collections.Generic.Dictionary _explicitEncryptionContext ; + private System.Collections.Generic.Dictionary _encryptionContext ; + private System.Collections.Generic.List _requiredContextFieldsOnDecrypt ; public AWS.Cryptography.StructuredEncryption.StructuredData PlaintextStructure { get { return this._plaintextStructure; } set { this._plaintextStructure = value; } @@ -38,19 +38,19 @@ public AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager Cmm { internal bool IsSetCmm () { return this._cmm != null; } - public System.Collections.Generic.Dictionary ImplicitEncryptionContext { - get { return this._implicitEncryptionContext; } - set { this._implicitEncryptionContext = value; } + public System.Collections.Generic.Dictionary EncryptionContext { + get { return this._encryptionContext; } + set { this._encryptionContext = value; } } - internal bool IsSetImplicitEncryptionContext () { - return this._implicitEncryptionContext != null; + internal bool IsSetEncryptionContext () { + return this._encryptionContext != null; } - public System.Collections.Generic.Dictionary ExplicitEncryptionContext { - get { return this._explicitEncryptionContext; } - set { this._explicitEncryptionContext = value; } + public System.Collections.Generic.List RequiredContextFieldsOnDecrypt { + get { return this._requiredContextFieldsOnDecrypt; } + set { this._requiredContextFieldsOnDecrypt = value; } } - internal bool IsSetExplicitEncryptionContext () { - return this._explicitEncryptionContext != null; + internal bool IsSetRequiredContextFieldsOnDecrypt () { + return this._requiredContextFieldsOnDecrypt != null; } public void Validate() { if (!IsSetPlaintextStructure()) throw new System.ArgumentException("Missing value for required property 'PlaintextStructure'"); diff --git a/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs b/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs index 8a34be042..86ccd8d13 100644 --- a/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs +++ b/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs @@ -3,31 +3,17 @@ // Do not modify this file. This file is machine generated, and any changes to it will be overwritten. using System.Linq; namespace AWS.Cryptography.StructuredEncryption { internal static class TypeConversion { - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_implicitEncryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M17_encryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); } - public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_implicitEncryptionContext (System.Collections.Generic.Dictionary value) { + public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M17_encryptionContext (System.Collections.Generic.Dictionary value) { return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); } - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_explicitEncryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { - return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); -} - public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_explicitEncryptionContext (System.Collections.Generic.Dictionary value) { - return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key (Dafny.ISequence value) { - return FromDafny_N6_smithy__N3_api__S6_String(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key (string value) { - return ToDafny_N6_smithy__N3_api__S6_String(value); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext (Dafny.IMap, Dafny.ISequence> value) { - return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value(pair.Cdr)); + public static System.Collections.Generic.List FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M30_requiredContextFieldsOnDecrypt (Wrappers_Compile._IOption>> value) { + return value.is_None ? (System.Collections.Generic.List) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList(value.Extract()); } - public static Dafny.IMap, Dafny.ISequence> ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext (System.Collections.Generic.Dictionary value) { - return Dafny.Map, Dafny.ISequence>.FromCollection(value.Select(pair => - new Dafny.Pair, Dafny.ISequence>(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value(pair.Value)) -)); + public static Wrappers_Compile._IOption>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M30_requiredContextFieldsOnDecrypt (System.Collections.Generic.List value) { + return value == null ? Wrappers_Compile.Option>>.create_None() : Wrappers_Compile.Option>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList((System.Collections.Generic.List) value)); } public static string FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes (Dafny.ISequence value) { System.Text.UTF8Encoding utf8 = new System.Text.UTF8Encoding(false, true); @@ -47,6 +33,174 @@ public static Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Structured return new Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_StructuredEncryptionException ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message(value.Message) ) ; +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes (Wrappers_Compile._IOption, Dafny.ISequence>> value) { + return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes(value.Extract()); +} + public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes (System.Collections.Generic.Dictionary value) { + return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes((System.Collections.Generic.Dictionary) value)); +} + public static string FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value (Dafny.ISequence value) { + return FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value (string value) { + return ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes (Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction> value) { + return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value(pair.Cdr)); +} + public static Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes (System.Collections.Generic.Dictionary value) { + return Dafny.Map, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>.FromCollection(value.Select(pair => + new Dafny.Pair, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value(pair.Value)) +)); +} + public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData)value; AWS.Cryptography.StructuredEncryption.StructuredData converted = new AWS.Cryptography.StructuredEncryption.StructuredData(); converted.Content = (AWS.Cryptography.StructuredEncryption.StructuredDataContent) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content(concrete.content); + if (concrete.attributes.is_Some) converted.Attributes = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes(concrete.attributes); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData (AWS.Cryptography.StructuredEncryption.StructuredData value) { + System.Collections.Generic.Dictionary var_attributes = value.IsSetAttributes() ? value.Attributes : (System.Collections.Generic.Dictionary) null; + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content(value.Content) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes(var_attributes) ) ; +} + public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static AWS.Cryptography.StructuredEncryption.CryptoSchemaContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent value) { + +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent value) { + +} + public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList__M6_member (Dafny.ISequence value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S9_Utf8Bytes(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList__M6_member (string value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S9_Utf8Bytes(value); +} + public static AWS.Cryptography.StructuredEncryption.StructuredDataContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content (AWS.Cryptography.StructuredEncryption.StructuredDataContent value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent(value); +} + public static AWS.Cryptography.StructuredEncryption.CryptoAction FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction value) { + if (value.is_ENCRYPT__AND__SIGN) return AWS.Cryptography.StructuredEncryption.CryptoAction.ENCRYPT_AND_SIGN; + if (value.is_SIGN__ONLY) return AWS.Cryptography.StructuredEncryption.CryptoAction.SIGN_ONLY; + if (value.is_IGNORE) return AWS.Cryptography.StructuredEncryption.CryptoAction.IGNORE; +throw new System.ArgumentException("Invalid AWS.Cryptography.StructuredEncryption.CryptoAction value"); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction (AWS.Cryptography.StructuredEncryption.CryptoAction value) { + if (AWS.Cryptography.StructuredEncryption.CryptoAction.ENCRYPT_AND_SIGN.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_ENCRYPT__AND__SIGN(); + if (AWS.Cryptography.StructuredEncryption.CryptoAction.SIGN_ONLY.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_SIGN__ONLY(); + if (AWS.Cryptography.StructuredEncryption.CryptoAction.IGNORE.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_IGNORE(); +throw new System.ArgumentException("Invalid AWS.Cryptography.StructuredEncryption.CryptoAction value"); +} + public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); +} + public static AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_StructuredEncryptionConfig (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredEncryptionConfig value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig)value; AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig converted = new AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig(); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredEncryptionConfig ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_StructuredEncryptionConfig (AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig value) { + + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig ( ) ; +} + public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { + return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { + return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M17_encryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { + return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); +} + public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M17_encryptionContext (System.Collections.Generic.Dictionary value) { + return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); +} + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key (Dafny.ISequence value) { + return FromDafny_N6_smithy__N3_api__S6_String(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key (string value) { + return ToDafny_N6_smithy__N3_api__S6_String(value); +} + public static AWS.Cryptography.StructuredEncryption.EncryptStructureInput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput)value; AWS.Cryptography.StructuredEncryption.EncryptStructureInput converted = new AWS.Cryptography.StructuredEncryption.EncryptStructureInput(); converted.PlaintextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(concrete.plaintextStructure); + converted.CryptoSchema = (AWS.Cryptography.StructuredEncryption.CryptoSchema) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(concrete.cryptoSchema); + if (concrete.keyring.is_Some) converted.Keyring = (AWS.Cryptography.MaterialProviders.IKeyring) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring(concrete.keyring); + if (concrete.cmm.is_Some) converted.Cmm = (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm(concrete.cmm); + if (concrete.encryptionContext.is_Some) converted.EncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M17_encryptionContext(concrete.encryptionContext); + if (concrete.requiredContextFieldsOnDecrypt.is_Some) converted.RequiredContextFieldsOnDecrypt = (System.Collections.Generic.List) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M30_requiredContextFieldsOnDecrypt(concrete.requiredContextFieldsOnDecrypt); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput (AWS.Cryptography.StructuredEncryption.EncryptStructureInput value) { + AWS.Cryptography.MaterialProviders.IKeyring var_keyring = value.IsSetKeyring() ? value.Keyring : (AWS.Cryptography.MaterialProviders.IKeyring) null; + AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager var_cmm = value.IsSetCmm() ? value.Cmm : (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null; + System.Collections.Generic.Dictionary var_encryptionContext = value.IsSetEncryptionContext() ? value.EncryptionContext : (System.Collections.Generic.Dictionary) null; + System.Collections.Generic.List var_requiredContextFieldsOnDecrypt = value.IsSetRequiredContextFieldsOnDecrypt() ? value.RequiredContextFieldsOnDecrypt : (System.Collections.Generic.List) null; + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(value.PlaintextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(value.CryptoSchema) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring(var_keyring) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm(var_cmm) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M17_encryptionContext(var_encryptionContext) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M30_requiredContextFieldsOnDecrypt(var_requiredContextFieldsOnDecrypt) ) ; +} + public static AWS.Cryptography.MaterialProviders.IKeyring FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring (Wrappers_Compile._IOption value) { + return value.is_None ? (AWS.Cryptography.MaterialProviders.IKeyring) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference(value.Extract()); +} + public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring (AWS.Cryptography.MaterialProviders.IKeyring value) { + return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference((AWS.Cryptography.MaterialProviders.IKeyring) value)); +} + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion (Dafny.ISequence value) { + return new string(value.Elements); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion (string value) { + return Dafny.Sequence.FromString(value); +} + public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm (Wrappers_Compile._IOption value) { + return value.is_None ? (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference(value.Extract()); +} + public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager value) { + return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference((AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) value)); +} + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message (Dafny.ISequence value) { + return FromDafny_N6_smithy__N3_api__S6_String(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message (string value) { + return ToDafny_N6_smithy__N3_api__S6_String(value); +} + public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { + Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema)value; AWS.Cryptography.StructuredEncryption.CryptoSchema converted = new AWS.Cryptography.StructuredEncryption.CryptoSchema(); converted.Content = (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content(concrete.content); + if (concrete.attributes.is_Some) converted.Attributes = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes(concrete.attributes); return converted; +} + public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { + System.Collections.Generic.Dictionary var_attributes = value.IsSetAttributes() ? value.Attributes : (System.Collections.Generic.Dictionary) null; + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content(value.Content) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes(var_attributes) ) ; +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes (Wrappers_Compile._IOption, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>> value) { + return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes(value.Extract()); +} + public static Wrappers_Compile._IOption, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes (System.Collections.Generic.Dictionary value) { + return value == null ? Wrappers_Compile.Option, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>>.create_None() : Wrappers_Compile.Option, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes((System.Collections.Generic.Dictionary) value)); +} + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key (Dafny.ISequence value) { + return FromDafny_N6_smithy__N3_api__S6_String(value); +} + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key (string value) { + return ToDafny_N6_smithy__N3_api__S6_String(value); +} + public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext (Dafny.IMap, Dafny.ISequence> value) { + return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value(pair.Cdr)); +} + public static Dafny.IMap, Dafny.ISequence> ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext (System.Collections.Generic.Dictionary value) { + return Dafny.Map, Dafny.ISequence>.FromCollection(value.Select(pair => + new Dafny.Pair, Dafny.ISequence>(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value(pair.Value)) +)); } public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference (Dafny.Aws.Cryptography.MaterialProviders.Types.ICryptographicMaterialsManager value) { if (value is NativeWrapper_CryptographicMaterialsManager nativeWrapper) return nativeWrapper._impl; @@ -95,70 +249,24 @@ public static string FromDafny_N6_smithy__N3_api__S6_String (Dafny.ISequence ToDafny_N6_smithy__N3_api__S6_String (string value) { return Dafny.Sequence.FromString(value); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes (Wrappers_Compile._IOption, Dafny.ISequence>> value) { - return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes(value.Extract()); -} - public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes (System.Collections.Generic.Dictionary value) { - return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes((System.Collections.Generic.Dictionary) value)); -} - public static string FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value (Dafny.ISequence value) { - return FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value (string value) { - return ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); } public static System.IO.MemoryStream FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal (Dafny.ISequence value) { return new System.IO.MemoryStream(value.Elements); } public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal (System.IO.MemoryStream value) { return Dafny.Sequence.FromArray(value.ToArray()); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes (Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction> value) { - return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value(pair.Cdr)); -} - public static Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes (System.Collections.Generic.Dictionary value) { - return Dafny.Map, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>.FromCollection(value.Select(pair => - new Dafny.Pair, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value(pair.Value)) -)); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_explicitEncryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { - return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); -} - public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_explicitEncryptionContext (System.Collections.Generic.Dictionary value) { - return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); -} - public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData)value; AWS.Cryptography.StructuredEncryption.StructuredData converted = new AWS.Cryptography.StructuredEncryption.StructuredData(); converted.Content = (AWS.Cryptography.StructuredEncryption.StructuredDataContent) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content(concrete.content); - if (concrete.attributes.is_Some) converted.Attributes = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes(concrete.attributes); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData (AWS.Cryptography.StructuredEncryption.StructuredData value) { - System.Collections.Generic.Dictionary var_attributes = value.IsSetAttributes() ? value.Attributes : (System.Collections.Generic.Dictionary) null; - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content(value.Content) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes(var_attributes) ) ; } public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key (Dafny.ISequence value) { return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion(value); } public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key (string value) { return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion(value); -} - public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); } public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); } public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); -} - public static AWS.Cryptography.StructuredEncryption.CryptoSchemaContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent value) { - -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent value) { - } public static AWS.Cryptography.MaterialProviders.IKeyring FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference (Dafny.Aws.Cryptography.MaterialProviders.Types.IKeyring value) { if (value is NativeWrapper_Keyring nativeWrapper) return nativeWrapper._impl; @@ -176,12 +284,6 @@ public static Dafny.Aws.Cryptography.MaterialProviders.Types.IKeyring ToDafny_N3 throw new System.ArgumentException( "Custom implementations of Keyring must extend KeyringBase."); } -} - public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); } public static AWS.Cryptography.StructuredEncryption.CryptoAction FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction value) { return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction(value); @@ -201,24 +303,6 @@ public static string FromDafny_N3_aws__N12_cryptography__N17_materialProviders__ } public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key (string value) { return ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); -} - public static AWS.Cryptography.StructuredEncryption.StructuredDataContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content (AWS.Cryptography.StructuredEncryption.StructuredDataContent value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent(value); -} - public static AWS.Cryptography.StructuredEncryption.CryptoAction FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction value) { - if (value.is_ENCRYPT__AND__SIGN) return AWS.Cryptography.StructuredEncryption.CryptoAction.ENCRYPT_AND_SIGN; - if (value.is_SIGN__ONLY) return AWS.Cryptography.StructuredEncryption.CryptoAction.SIGN_ONLY; - if (value.is_IGNORE) return AWS.Cryptography.StructuredEncryption.CryptoAction.IGNORE; -throw new System.ArgumentException("Invalid AWS.Cryptography.StructuredEncryption.CryptoAction value"); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction (AWS.Cryptography.StructuredEncryption.CryptoAction value) { - if (AWS.Cryptography.StructuredEncryption.CryptoAction.ENCRYPT_AND_SIGN.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_ENCRYPT__AND__SIGN(); - if (AWS.Cryptography.StructuredEncryption.CryptoAction.SIGN_ONLY.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_SIGN__ONLY(); - if (AWS.Cryptography.StructuredEncryption.CryptoAction.IGNORE.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_IGNORE(); -throw new System.ArgumentException("Invalid AWS.Cryptography.StructuredEncryption.CryptoAction value"); } public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes (Dafny.IMap, Dafny.ISequence> value) { return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value(pair.Cdr)); @@ -227,31 +311,6 @@ public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction T return Dafny.Map, Dafny.ISequence>.FromCollection(value.Select(pair => new Dafny.Pair, Dafny.ISequence>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value(pair.Value)) )); -} - public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); -} - public static AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_StructuredEncryptionConfig (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredEncryptionConfig value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig)value; AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig converted = new AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig(); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredEncryptionConfig ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_StructuredEncryptionConfig (AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig value) { - - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig ( ) ; -} - public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key (Dafny.ISequence value) { - return FromDafny_N6_smithy__N3_api__S6_String(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key (string value) { - return ToDafny_N6_smithy__N3_api__S6_String(value); } public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm (Wrappers_Compile._IOption value) { return value.is_None ? (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference(value.Extract()); @@ -259,32 +318,11 @@ public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager value) { return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference((AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) value)); } - public static AWS.Cryptography.StructuredEncryption.EncryptStructureInput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput)value; AWS.Cryptography.StructuredEncryption.EncryptStructureInput converted = new AWS.Cryptography.StructuredEncryption.EncryptStructureInput(); converted.PlaintextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(concrete.plaintextStructure); - converted.CryptoSchema = (AWS.Cryptography.StructuredEncryption.CryptoSchema) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(concrete.cryptoSchema); - if (concrete.keyring.is_Some) converted.Keyring = (AWS.Cryptography.MaterialProviders.IKeyring) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring(concrete.keyring); - if (concrete.cmm.is_Some) converted.Cmm = (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm(concrete.cmm); - if (concrete.implicitEncryptionContext.is_Some) converted.ImplicitEncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_implicitEncryptionContext(concrete.implicitEncryptionContext); - if (concrete.explicitEncryptionContext.is_Some) converted.ExplicitEncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_explicitEncryptionContext(concrete.explicitEncryptionContext); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput (AWS.Cryptography.StructuredEncryption.EncryptStructureInput value) { - AWS.Cryptography.MaterialProviders.IKeyring var_keyring = value.IsSetKeyring() ? value.Keyring : (AWS.Cryptography.MaterialProviders.IKeyring) null; - AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager var_cmm = value.IsSetCmm() ? value.Cmm : (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null; - System.Collections.Generic.Dictionary var_implicitEncryptionContext = value.IsSetImplicitEncryptionContext() ? value.ImplicitEncryptionContext : (System.Collections.Generic.Dictionary) null; - System.Collections.Generic.Dictionary var_explicitEncryptionContext = value.IsSetExplicitEncryptionContext() ? value.ExplicitEncryptionContext : (System.Collections.Generic.Dictionary) null; - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(value.PlaintextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(value.CryptoSchema) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring(var_keyring) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm(var_cmm) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_implicitEncryptionContext(var_implicitEncryptionContext) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_explicitEncryptionContext(var_explicitEncryptionContext) ) ; + public static System.Collections.Generic.List FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList (Dafny.ISequence> value) { + return new System.Collections.Generic.List(value.Elements.Select(FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList__M6_member)); } - public static AWS.Cryptography.MaterialProviders.IKeyring FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring (Wrappers_Compile._IOption value) { - return value.is_None ? (AWS.Cryptography.MaterialProviders.IKeyring) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference(value.Extract()); -} - public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring (AWS.Cryptography.MaterialProviders.IKeyring value) { - return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference((AWS.Cryptography.MaterialProviders.IKeyring) value)); -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion (Dafny.ISequence value) { - return new string(value.Elements); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion (string value) { - return Dafny.Sequence.FromString(value); + public static Dafny.ISequence> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList (System.Collections.Generic.List value) { + return Dafny.Sequence>.FromArray(value.Select(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList__M6_member).ToArray()); } public static System.IO.MemoryStream FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value (Dafny.ISequence value) { return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal(value); @@ -297,47 +335,27 @@ public static AWS.Cryptography.StructuredEncryption.DecryptStructureInput FromDa converted.CryptoSchemas = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas(concrete.cryptoSchemas); if (concrete.keyring.is_Some) converted.Keyring = (AWS.Cryptography.MaterialProviders.IKeyring) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring(concrete.keyring); if (concrete.cmm.is_Some) converted.Cmm = (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm(concrete.cmm); - if (concrete.implicitEncryptionContext.is_Some) converted.ImplicitEncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_implicitEncryptionContext(concrete.implicitEncryptionContext); - if (concrete.explicitEncryptionContext.is_Some) converted.ExplicitEncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_explicitEncryptionContext(concrete.explicitEncryptionContext); return converted; + if (concrete.encryptionContext.is_Some) converted.EncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M17_encryptionContext(concrete.encryptionContext); return converted; } public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureInput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput (AWS.Cryptography.StructuredEncryption.DecryptStructureInput value) { AWS.Cryptography.MaterialProviders.IKeyring var_keyring = value.IsSetKeyring() ? value.Keyring : (AWS.Cryptography.MaterialProviders.IKeyring) null; AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager var_cmm = value.IsSetCmm() ? value.Cmm : (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null; - System.Collections.Generic.Dictionary var_implicitEncryptionContext = value.IsSetImplicitEncryptionContext() ? value.ImplicitEncryptionContext : (System.Collections.Generic.Dictionary) null; - System.Collections.Generic.Dictionary var_explicitEncryptionContext = value.IsSetExplicitEncryptionContext() ? value.ExplicitEncryptionContext : (System.Collections.Generic.Dictionary) null; - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure(value.CiphertextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas(value.CryptoSchemas) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring(var_keyring) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm(var_cmm) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_implicitEncryptionContext(var_implicitEncryptionContext) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M25_explicitEncryptionContext(var_explicitEncryptionContext) ) ; -} - public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm (Wrappers_Compile._IOption value) { - return value.is_None ? (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference(value.Extract()); + System.Collections.Generic.Dictionary var_encryptionContext = value.IsSetEncryptionContext() ? value.EncryptionContext : (System.Collections.Generic.Dictionary) null; + return new Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure(value.CiphertextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas(value.CryptoSchemas) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring(var_keyring) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm(var_cmm) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M17_encryptionContext(var_encryptionContext) ) ; } - public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager value) { - return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference((AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) value)); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_implicitEncryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { - return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); -} - public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M25_implicitEncryptionContext (System.Collections.Generic.Dictionary value) { - return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message (Dafny.ISequence value) { - return FromDafny_N6_smithy__N3_api__S6_String(value); + public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S9_Utf8Bytes (Dafny.ISequence value) { + System.Text.UTF8Encoding utf8 = new System.Text.UTF8Encoding(false, true); +return utf8.GetString(value.Elements); } - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message (string value) { - return ToDafny_N6_smithy__N3_api__S6_String(value); + public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S9_Utf8Bytes (string value) { + System.Text.UTF8Encoding utf8 = new System.Text.UTF8Encoding(false, true); +return Dafny.Sequence.FromArray(utf8.GetBytes(value)); } public static AWS.Cryptography.MaterialProviders.IKeyring FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring (Wrappers_Compile._IOption value) { return value.is_None ? (AWS.Cryptography.MaterialProviders.IKeyring) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference(value.Extract()); } public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring (AWS.Cryptography.MaterialProviders.IKeyring value) { return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference((AWS.Cryptography.MaterialProviders.IKeyring) value)); -} - public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema)value; AWS.Cryptography.StructuredEncryption.CryptoSchema converted = new AWS.Cryptography.StructuredEncryption.CryptoSchema(); converted.Content = (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content(concrete.content); - if (concrete.attributes.is_Some) converted.Attributes = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes(concrete.attributes); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { - System.Collections.Generic.Dictionary var_attributes = value.IsSetAttributes() ? value.Attributes : (System.Collections.Generic.Dictionary) null; - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content(value.Content) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes(var_attributes) ) ; } public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas (Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema> value) { return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value(pair.Cdr)); @@ -346,12 +364,6 @@ public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema T return Dafny.Map, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema>.FromCollection(value.Select(pair => new Dafny.Pair, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value(pair.Value)) )); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes (Wrappers_Compile._IOption, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>> value) { - return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes(value.Extract()); -} - public static Wrappers_Compile._IOption, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes (System.Collections.Generic.Dictionary value) { - return value == null ? Wrappers_Compile.Option, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>>.create_None() : Wrappers_Compile.Option, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes((System.Collections.Generic.Dictionary) value)); } public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); diff --git a/src/StructuredEncryption/src/Index.dfy b/src/StructuredEncryption/src/Index.dfy new file mode 100644 index 000000000..af843ca0f --- /dev/null +++ b/src/StructuredEncryption/src/Index.dfy @@ -0,0 +1,82 @@ +include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" +include "Operations/EncryptStructureOperation.dfy" +include "Operations/DecryptStructureOperation.dfy" +// TODO including UInt in this file causes build issues. It complains of duplicate UInt modules... +include "../Model/AwsCryptographyStructuredEncryptionTypes.dfy" + +module {:extern "Dafny.Aws.StructuredEncryption.StructuredEncryptionClient"} StructuredEncryptionClient { + import opened Wrappers + import opened StandardLibrary + import Types = AwsCryptographyStructuredEncryptionTypes + import EncryptStructureOperation + import DecryptStructureOperation + + class StructuredEncryptionClient extends Types.IStructuredEncryptionClient { + predicate ValidState() + ensures ValidState() ==> History in Modifies + { + History in Modifies + } + + constructor() + ensures ValidState() && fresh(Modifies) && fresh(History) + { + History := new Types.IStructuredEncryptionClientCallHistory(); + Modifies := {History}; + } + + predicate EncryptStructureEnsuresPublicly( + input: Types.EncryptStructureInput, + output: Result) + { + var encryptionContextFields := if input.encryptionContext.Some? then input.encryptionContext.value.Keys else {}; + var requiredFields := if input.requiredContextFieldsOnDecrypt.Some? then input.requiredContextFieldsOnDecrypt.value else []; + && !(forall k :: k in requiredFields ==> k in encryptionContextFields) ==> output.Failure? + && ( + || (input.cmm.Some? && input.keyring.Some?) + || (input.cmm.None? && input.keyring.None?) + ) ==> output.Failure? + } + + predicate DecryptStructureEnsuresPublicly( + input: Types.DecryptStructureInput, + output: Result) + { + && ( + || (input.cmm.Some? && input.keyring.Some?) + || (input.cmm.None? && input.keyring.None?) + ) ==> output.Failure? + && |input.cryptoSchemas| <= 0 ==> output.Failure? + } + + method EncryptStructure(input: Types.EncryptStructureInput) + returns (output: Result) + requires ValidState() + modifies Modifies - {History}, + History`EncryptStructure + decreases Modifies , + (if input.keyring.Some? then input.keyring.value.Modifies else {}) , + (if input.cmm.Some? then input.cmm.value.Modifies else {}) + ensures EncryptStructureEnsuresPublicly(input, output) + ensures History.EncryptStructure == old(History.EncryptStructure) + [Types.DafnyCallEvent(input, output)] + { + output := EncryptStructureOperation.EncryptStructure(input); + History.EncryptStructure := History.EncryptStructure + [Types.DafnyCallEvent(input, output)]; + } + + method DecryptStructure(input: Types.DecryptStructureInput) + returns (output: Result) + requires ValidState() + modifies Modifies - {History}, + History`DecryptStructure + decreases Modifies , + (if input.keyring.Some? then input.keyring.value.Modifies else {}) , + (if input.cmm.Some? then input.cmm.value.Modifies else {}) + ensures DecryptStructureEnsuresPublicly(input, output) + ensures History.DecryptStructure == old(History.DecryptStructure) + [Types.DafnyCallEvent(input, output)] + { + output := DecryptStructureOperation.DecryptStructure(input); + History.DecryptStructure := History.DecryptStructure + [Types.DafnyCallEvent(input, output)]; + } + } +} \ No newline at end of file diff --git a/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy b/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy new file mode 100644 index 000000000..90b6f79ca --- /dev/null +++ b/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy @@ -0,0 +1,56 @@ +include "../../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" +include "../../Model/AwsCryptographyStructuredEncryptionTypes.dfy" + +module DecryptStructureOperation { + import opened Wrappers + import opened StandardLibrary + import Types = AwsCryptographyStructuredEncryptionTypes + + method DecryptStructure(input: Types.DecryptStructureInput) + returns (output: Result) + ensures + ( + || (input.cmm.Some? && input.keyring.Some?) + || (input.cmm.None? && input.keyring.None?) + ) ==> output.Failure? + ensures |input.cryptoSchemas| <= 0 ==> output.Failure? + { + // Ensure a Keyring XOR a CMM was included on input, otherwise error + :- Need(input.cmm.Some? || input.keyring.Some?, + Types.Error.StructuredEncryptionException( + message := "Invalid input: A Keyring or CMM MUST be supplied on input.")); + :- Need(input.cmm.None? || input.keyring.None?, + Types.Error.StructuredEncryptionException( + message := "Invalid input: Cannot recieve both a Keyring and a CMM on input.")); + + // Ensure CryptoSchemas is not empty + :- Need(|input.cryptoSchemas| > 0, + Types.Error.StructuredEncryptionException( + message := "Invalid input: a non-empty Crypto Schema map must be supplied on input.")); + + // TODO: Currently stubbed out to return a hardcoded StructuredData + var stubbedBytes := [0x68, 0x65, 0x6c, 0x6c, 0x6f, 0x2c, 0x20, 0x77, 0x6f, 0x72, 0x6c, 0x64]; + var stubbedStructure := Types.StructuredData( + content := Types.StructuredDataContent.dataMap( + StructuredDataMap := map[ + "foo" := Types.StructuredData( + content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "bar" := Types.StructuredData( + content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "fizzbuzz" := Types.StructuredData( + content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ) + ] + ), + attributes := None() + ); + + var decryptOutput := Types.DecryptStructureOutput(plaintextStructure := stubbedStructure); + output := Success(decryptOutput); + } +} \ No newline at end of file diff --git a/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy b/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy new file mode 100644 index 000000000..6cc7c59f5 --- /dev/null +++ b/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy @@ -0,0 +1,61 @@ +include "../../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" +include "../../Model/AwsCryptographyStructuredEncryptionTypes.dfy" + +module EncryptStructureOperation { + import opened Wrappers + import opened StandardLibrary + import Types = AwsCryptographyStructuredEncryptionTypes + + method EncryptStructure(input: Types.EncryptStructureInput) + returns (output: Result) + ensures ( + || (input.cmm.Some? && input.keyring.Some?) + || (input.cmm.None? && input.keyring.None?) + ) ==> output.Failure? + ensures + var encryptionContextFields := if input.encryptionContext.Some? then input.encryptionContext.value.Keys else {}; + var requiredFields := if input.requiredContextFieldsOnDecrypt.Some? then input.requiredContextFieldsOnDecrypt.value else []; + !(forall k :: k in requiredFields ==> k in encryptionContextFields) ==> output.Failure? + { + // Ensure a Keyring XOR a CMM was included on input, otherwise error + :- Need(input.cmm.Some? || input.keyring.Some?, + Types.Error.StructuredEncryptionException( + message := "Invalid input: A Keyring or CMM MUST be supplied on input.")); + :- Need(input.cmm.None? || input.keyring.None?, + Types.Error.StructuredEncryptionException( + message := "Invalid input: Cannot recieve both a Keyring and a CMM on input.")); + + // TODO Once verified, this information should probably move to some new EC datatype that combines this info + // This datatype should live in the MPL. + :- Need(input.requiredContextFieldsOnDecrypt.Some? ==> ( + && input.encryptionContext.Some? + && forall k :: k in input.requiredContextFieldsOnDecrypt.value ==> k in input.encryptionContext.value.Keys), + Types.Error.StructuredEncryptionException( + message := "Invalid input: A Keyring or CMM MUST be supplied on input.")); + + // TODO: Currently stubbed out to return a hardcoded StructuredData + var stubbedBytes := [0x21, 0x64, 0x6c, 0x72, 0x6f, 0x77, 0x20, 0x2c, 0x6f, 0x6c, 0x6c, 0x65, 0x68]; + var stubbedStructure := Types.StructuredData( + content := Types.StructuredDataContent.dataMap( + StructuredDataMap := map[ + "foo" := Types.StructuredData( + content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "bar" := Types.StructuredData( + content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ), + "fizzbuzz" := Types.StructuredData( + content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), + attributes := None() + ) + ] + ), + attributes := None() + ); + + var encryptOutput := Types.EncryptStructureOutput(ciphertextStructure := stubbedStructure); + output := Success(encryptOutput); + } +} \ No newline at end of file diff --git a/src/StructuredEncryption/src/StructuredEncryptionClient.dfy b/src/StructuredEncryption/src/StructuredEncryptionClient.dfy deleted file mode 100644 index 31f2b3cc0..000000000 --- a/src/StructuredEncryption/src/StructuredEncryptionClient.dfy +++ /dev/null @@ -1,130 +0,0 @@ -include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" -// TODO including UInt in this file causes build issues. It complains of duplicate UInt modules... -include "../Model/AwsCryptographyStructuredEncryptionTypes.dfy" - -module {:extern "Dafny.Aws.StructuredEncryption.StructuredEncryptionClient"} StructuredEncryptionClient { - import opened Wrappers - import opened StandardLibrary - import Types = AwsCryptographyStructuredEncryptionTypes - - class StructuredEncryptionClient extends Types.IStructuredEncryptionClient { - predicate ValidState() - ensures ValidState() ==> History in Modifies - { - History in Modifies - } - - constructor() - ensures ValidState() && fresh(Modifies) && fresh(History) - { - History := new Types.IStructuredEncryptionClientCallHistory(); - Modifies := {History}; - } - - predicate EncryptStructureEnsuresPublicly( - input: Types.EncryptStructureInput, - output: Result) - { - true - } - - predicate DecryptStructureEnsuresPublicly( - input: Types.DecryptStructureInput, - output: Result) - { - true - } - - method EncryptStructure(input: Types.EncryptStructureInput) - returns (output: Result) - requires ValidState() - modifies Modifies - {History}, - History`EncryptStructure - ensures EncryptStructureEnsuresPublicly(input, output) - ensures History.EncryptStructure == old(History.EncryptStructure) + [Types.DafnyCallEvent(input, output)] - { - // Ensure a Keyring XOR a CMM was included on input, otherwise error - :- Need(input.cmm.Some? || input.keyring.Some?, - Types.Error.StructuredEncryptionException( - message := "Invalid input: A Keyring or CMM MUST be supplied on input.")); - :- Need(input.cmm.None? || input.keyring.None?, - Types.Error.StructuredEncryptionException( - message := "Invalid input: Cannot recieve both a Keyring and a CMM on input.")); - - // TODO: Currently stubbed out to return a hardcoded StructuredData - var stubbedBytes := [0x21, 0x64, 0x6c, 0x72, 0x6f, 0x77, 0x20, 0x2c, 0x6f, 0x6c, 0x6c, 0x65, 0x68]; - var stubbedStructure := Types.StructuredData( - content := Types.StructuredDataContent.dataMap( - StructuredDataMap := map[ - "foo" := Types.StructuredData( - content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), - attributes := None() - ), - "bar" := Types.StructuredData( - content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), - attributes := None() - ), - "fizzbuzz" := Types.StructuredData( - content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), - attributes := None() - ) - ] - ), - attributes := None() - ); - - var encryptOutput := Types.EncryptStructureOutput(ciphertextStructure := stubbedStructure); - - output := Success(encryptOutput); - History.EncryptStructure := History.EncryptStructure + [Types.DafnyCallEvent(input, output)]; - } - - method DecryptStructure(input: Types.DecryptStructureInput) - returns (output: Result) - requires ValidState() - modifies Modifies - {History}, - History`DecryptStructure - ensures DecryptStructureEnsuresPublicly(input, output) - ensures History.DecryptStructure == old(History.DecryptStructure) + [Types.DafnyCallEvent(input, output)] - { - // Ensure a Keyring XOR a CMM was included on input, otherwise error - :- Need(input.cmm.Some? || input.keyring.Some?, - Types.Error.StructuredEncryptionException( - message := "Invalid input: A Keyring or CMM MUST be supplied on input.")); - :- Need(input.cmm.None? || input.keyring.None?, - Types.Error.StructuredEncryptionException( - message := "Invalid input: Cannot recieve both a Keyring and a CMM on input.")); - - // Ensure CryptoSchemas is not empty - :- Need(|input.cryptoSchemas| > 0, - Types.Error.StructuredEncryptionException( - message := "Invalid input: a non-empty Crypto Schema map must be supplied on input.")); - - // TODO: Currently stubbed out to return a hardcoded StructuredData - var stubbedBytes := [0x68, 0x65, 0x6c, 0x6c, 0x6f, 0x2c, 0x20, 0x77, 0x6f, 0x72, 0x6c, 0x64]; - var stubbedStructure := Types.StructuredData( - content := Types.StructuredDataContent.dataMap( - StructuredDataMap := map[ - "foo" := Types.StructuredData( - content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), - attributes := None() - ), - "bar" := Types.StructuredData( - content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), - attributes := None() - ), - "fizzbuzz" := Types.StructuredData( - content := Types.StructuredDataContent.terminal(Terminal:=stubbedBytes), - attributes := None() - ) - ] - ), - attributes := None() - ); - - var decryptOutput := Types.DecryptStructureOutput(plaintextStructure := stubbedStructure); - output := Success(decryptOutput); - History.DecryptStructure := History.DecryptStructure + [Types.DafnyCallEvent(input, output)]; - } - } -} \ No newline at end of file diff --git a/src/StructuredEncryption/test/HappyCaseTests.dfy b/src/StructuredEncryption/test/HappyCaseTests.dfy index 5713ea27b..89bd140ac 100644 --- a/src/StructuredEncryption/test/HappyCaseTests.dfy +++ b/src/StructuredEncryption/test/HappyCaseTests.dfy @@ -1,5 +1,5 @@ include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" -include "../src/StructuredEncryptionClient.dfy" +include "../src/Index.dfy" include "../Model/AwsCryptographyStructuredEncryptionTypes.dfy" include "../../../private-aws-encryption-sdk-dafny-staging/src/AwsCryptographicMaterialProviders/src/Index.dfy" @@ -45,8 +45,8 @@ module HappyCaseTests { cryptoSchema:=schema, keyring:=Some(keyring), cmm:=None(), - implicitEncryptionContext:=None(), - explicitEncryptionContext:=None() + encryptionContext:=None(), + requiredContextFieldsOnDecrypt:=None() ) ); @@ -111,8 +111,7 @@ module HappyCaseTests { cryptoSchemas:=schemaMap, keyring:=Some(keyring), cmm:=None(), - implicitEncryptionContext:=None(), - explicitEncryptionContext:=None() + encryptionContext:=None() ) ); From e9fa7a404328adee194ec873f717391e165aaf9b Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Mon, 22 Aug 2022 15:05:04 -0700 Subject: [PATCH 08/13] Clean up build --- Makefile | 8 +++++--- .../Test/AwsStructuredEncryptionTest.csproj | 20 ++----------------- .../src/StructuredEncryptionDafny.csproj | 8 ++------ 3 files changed, 9 insertions(+), 27 deletions(-) diff --git a/Makefile b/Makefile index b254cb7bd..42a3b8c3c 100644 --- a/Makefile +++ b/Makefile @@ -9,13 +9,15 @@ build: build-dotnet test: test-dotnet +# TODO cleanup CS0618 warning in MPL build-dotnet: cd src/StructuredEncryption/src;\ - dotnet build + dotnet build /nowarn:CS0618 -test-dotnet: build +# TODO cleanup CS0618 warning in MPL +test-dotnet: cd src/StructuredEncryption/runtimes/net;\ - dotnet test -f netcoreapp3.1 Test + dotnet test -f netcoreapp3.1 Test /nowarn:CS0618 generate-models: cd polymorph/smithy-dotnet;\ diff --git a/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj b/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj index 998179110..f0533f1d7 100644 --- a/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj +++ b/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj @@ -50,19 +50,15 @@ - - - - - - + + @@ -72,16 +68,4 @@ - - - - - - - - \ No newline at end of file diff --git a/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj b/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj index 91e26cbe6..8f35e2cb5 100644 --- a/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj +++ b/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj @@ -43,19 +43,15 @@ - - - - - - + + From 8b58b45ee4ef95f27d3273098414499c95fb522d Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Mon, 22 Aug 2022 16:46:38 -0700 Subject: [PATCH 09/13] cleanup --- .../model/StructuredEncryption.smithy | 1 - .../runtimes/net/Generated/CryptoAction.cs | 19 - .../runtimes/net/Generated/CryptoSchema.cs | 28 -- .../net/Generated/DecryptStructureInput.cs | 53 --- .../net/Generated/DecryptStructureOutput.cs | 20 - .../net/Generated/EncryptStructureInput.cs | 61 --- .../net/Generated/EncryptStructureOutput.cs | 20 - .../runtimes/net/Generated/OpaqueError.cs | 13 - .../runtimes/net/Generated/StructuredData.cs | 28 -- .../net/Generated/StructuredEncryption.cs | 31 -- .../Generated/StructuredEncryptionConfig.cs | 13 - .../StructuredEncryptionException.cs | 9 - .../runtimes/net/Generated/TypeConversion.cs | 402 ------------------ .../Test/AwsStructuredEncryptionTest.csproj | 8 +- src/StructuredEncryption/src/Index.dfy | 5 +- .../Operations/DecryptStructureOperation.dfy | 2 + .../Operations/EncryptStructureOperation.dfy | 2 + .../src/StructuredEncryptionDafny.csproj | 8 +- .../test/HappyCaseTests.dfy | 2 + 19 files changed, 17 insertions(+), 708 deletions(-) delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/CryptoAction.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/CryptoSchema.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/DecryptStructureOutput.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/EncryptStructureOutput.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/OpaqueError.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/StructuredData.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/StructuredEncryption.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionConfig.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionException.cs delete mode 100644 src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs diff --git a/src/StructuredEncryption/model/StructuredEncryption.smithy b/src/StructuredEncryption/model/StructuredEncryption.smithy index ec1c739c4..6798b0e67 100644 --- a/src/StructuredEncryption/model/StructuredEncryption.smithy +++ b/src/StructuredEncryption/model/StructuredEncryption.smithy @@ -12,7 +12,6 @@ use aws.polymorph#dafnyUtf8Bytes @dafnyUtf8Bytes string Utf8Bytes -// TODO: Bikeshed on name "StructuredEncryption" @aws.polymorph#localService( sdkId: "StructuredEncryption", config: StructuredEncryptionConfig, diff --git a/src/StructuredEncryption/runtimes/net/Generated/CryptoAction.cs b/src/StructuredEncryption/runtimes/net/Generated/CryptoAction.cs deleted file mode 100644 index ab58f2a96..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/CryptoAction.cs +++ /dev/null @@ -1,19 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { - using Amazon.Runtime; public class CryptoAction : ConstantClass { - - - public static readonly CryptoAction ENCRYPT_AND_SIGN = new CryptoAction ("ENCRYPT_AND_SIGN"); - - public static readonly CryptoAction SIGN_ONLY = new CryptoAction ("SIGN_ONLY"); - - public static readonly CryptoAction IGNORE = new CryptoAction ("IGNORE"); - public static readonly CryptoAction [] Values = { - ENCRYPT_AND_SIGN , IGNORE , SIGN_ONLY -} ; - public CryptoAction (string value) : base(value) {} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/CryptoSchema.cs b/src/StructuredEncryption/runtimes/net/Generated/CryptoSchema.cs deleted file mode 100644 index ed7206998..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/CryptoSchema.cs +++ /dev/null @@ -1,28 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { - public class CryptoSchema { - private AWS.Cryptography.StructuredEncryption.CryptoSchemaContent _content ; - private System.Collections.Generic.Dictionary _attributes ; - public AWS.Cryptography.StructuredEncryption.CryptoSchemaContent Content { - get { return this._content; } - set { this._content = value; } -} - internal bool IsSetContent () { - return this._content != null; -} - public System.Collections.Generic.Dictionary Attributes { - get { return this._attributes; } - set { this._attributes = value; } -} - internal bool IsSetAttributes () { - return this._attributes != null; -} - public void Validate() { - if (!IsSetContent()) throw new System.ArgumentException("Missing value for required property 'Content'"); - -} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs b/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs deleted file mode 100644 index ad1e7cfaf..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureInput.cs +++ /dev/null @@ -1,53 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { - public class DecryptStructureInput { - private AWS.Cryptography.StructuredEncryption.StructuredData _ciphertextStructure ; - private System.Collections.Generic.Dictionary _cryptoSchemas ; - private AWS.Cryptography.MaterialProviders.IKeyring _keyring ; - private AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager _cmm ; - private System.Collections.Generic.Dictionary _encryptionContext ; - public AWS.Cryptography.StructuredEncryption.StructuredData CiphertextStructure { - get { return this._ciphertextStructure; } - set { this._ciphertextStructure = value; } -} - internal bool IsSetCiphertextStructure () { - return this._ciphertextStructure != null; -} - public System.Collections.Generic.Dictionary CryptoSchemas { - get { return this._cryptoSchemas; } - set { this._cryptoSchemas = value; } -} - internal bool IsSetCryptoSchemas () { - return this._cryptoSchemas != null; -} - public AWS.Cryptography.MaterialProviders.IKeyring Keyring { - get { return this._keyring; } - set { this._keyring = value; } -} - internal bool IsSetKeyring () { - return this._keyring != null; -} - public AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager Cmm { - get { return this._cmm; } - set { this._cmm = value; } -} - internal bool IsSetCmm () { - return this._cmm != null; -} - public System.Collections.Generic.Dictionary EncryptionContext { - get { return this._encryptionContext; } - set { this._encryptionContext = value; } -} - internal bool IsSetEncryptionContext () { - return this._encryptionContext != null; -} - public void Validate() { - if (!IsSetCiphertextStructure()) throw new System.ArgumentException("Missing value for required property 'CiphertextStructure'"); - if (!IsSetCryptoSchemas()) throw new System.ArgumentException("Missing value for required property 'CryptoSchemas'"); - -} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureOutput.cs b/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureOutput.cs deleted file mode 100644 index 32a637f94..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/DecryptStructureOutput.cs +++ /dev/null @@ -1,20 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { - public class DecryptStructureOutput { - private AWS.Cryptography.StructuredEncryption.StructuredData _plaintextStructure ; - public AWS.Cryptography.StructuredEncryption.StructuredData PlaintextStructure { - get { return this._plaintextStructure; } - set { this._plaintextStructure = value; } -} - internal bool IsSetPlaintextStructure () { - return this._plaintextStructure != null; -} - public void Validate() { - if (!IsSetPlaintextStructure()) throw new System.ArgumentException("Missing value for required property 'PlaintextStructure'"); - -} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs b/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs deleted file mode 100644 index a230d4799..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureInput.cs +++ /dev/null @@ -1,61 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { - public class EncryptStructureInput { - private AWS.Cryptography.StructuredEncryption.StructuredData _plaintextStructure ; - private AWS.Cryptography.StructuredEncryption.CryptoSchema _cryptoSchema ; - private AWS.Cryptography.MaterialProviders.IKeyring _keyring ; - private AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager _cmm ; - private System.Collections.Generic.Dictionary _encryptionContext ; - private System.Collections.Generic.List _requiredContextFieldsOnDecrypt ; - public AWS.Cryptography.StructuredEncryption.StructuredData PlaintextStructure { - get { return this._plaintextStructure; } - set { this._plaintextStructure = value; } -} - internal bool IsSetPlaintextStructure () { - return this._plaintextStructure != null; -} - public AWS.Cryptography.StructuredEncryption.CryptoSchema CryptoSchema { - get { return this._cryptoSchema; } - set { this._cryptoSchema = value; } -} - internal bool IsSetCryptoSchema () { - return this._cryptoSchema != null; -} - public AWS.Cryptography.MaterialProviders.IKeyring Keyring { - get { return this._keyring; } - set { this._keyring = value; } -} - internal bool IsSetKeyring () { - return this._keyring != null; -} - public AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager Cmm { - get { return this._cmm; } - set { this._cmm = value; } -} - internal bool IsSetCmm () { - return this._cmm != null; -} - public System.Collections.Generic.Dictionary EncryptionContext { - get { return this._encryptionContext; } - set { this._encryptionContext = value; } -} - internal bool IsSetEncryptionContext () { - return this._encryptionContext != null; -} - public System.Collections.Generic.List RequiredContextFieldsOnDecrypt { - get { return this._requiredContextFieldsOnDecrypt; } - set { this._requiredContextFieldsOnDecrypt = value; } -} - internal bool IsSetRequiredContextFieldsOnDecrypt () { - return this._requiredContextFieldsOnDecrypt != null; -} - public void Validate() { - if (!IsSetPlaintextStructure()) throw new System.ArgumentException("Missing value for required property 'PlaintextStructure'"); - if (!IsSetCryptoSchema()) throw new System.ArgumentException("Missing value for required property 'CryptoSchema'"); - -} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureOutput.cs b/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureOutput.cs deleted file mode 100644 index a1b4af01f..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/EncryptStructureOutput.cs +++ /dev/null @@ -1,20 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { - public class EncryptStructureOutput { - private AWS.Cryptography.StructuredEncryption.StructuredData _ciphertextStructure ; - public AWS.Cryptography.StructuredEncryption.StructuredData CiphertextStructure { - get { return this._ciphertextStructure; } - set { this._ciphertextStructure = value; } -} - internal bool IsSetCiphertextStructure () { - return this._ciphertextStructure != null; -} - public void Validate() { - if (!IsSetCiphertextStructure()) throw new System.ArgumentException("Missing value for required property 'CiphertextStructure'"); - -} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/OpaqueError.cs b/src/StructuredEncryption/runtimes/net/Generated/OpaqueError.cs deleted file mode 100644 index be923ab7f..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/OpaqueError.cs +++ /dev/null @@ -1,13 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { - public class OpaqueError : Exception { - public readonly object obj; - public OpaqueError(Exception ex) : base("OpaqueError:", ex) { this.obj = ex; } - public OpaqueError() : base("Unknown Unexpected Error") { } - public OpaqueError(object obj) : base("Opaque obj is not an Exception.") { this.obj = obj;} -} - -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/StructuredData.cs b/src/StructuredEncryption/runtimes/net/Generated/StructuredData.cs deleted file mode 100644 index cc22c5b60..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/StructuredData.cs +++ /dev/null @@ -1,28 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { - public class StructuredData { - private AWS.Cryptography.StructuredEncryption.StructuredDataContent _content ; - private System.Collections.Generic.Dictionary _attributes ; - public AWS.Cryptography.StructuredEncryption.StructuredDataContent Content { - get { return this._content; } - set { this._content = value; } -} - internal bool IsSetContent () { - return this._content != null; -} - public System.Collections.Generic.Dictionary Attributes { - get { return this._attributes; } - set { this._attributes = value; } -} - internal bool IsSetAttributes () { - return this._attributes != null; -} - public void Validate() { - if (!IsSetContent()) throw new System.ArgumentException("Missing value for required property 'Content'"); - -} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryption.cs b/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryption.cs deleted file mode 100644 index d9f6a30ec..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryption.cs +++ /dev/null @@ -1,31 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using System.IO; - using System.Collections.Generic; - using AWS.Cryptography.StructuredEncryption; - using Dafny.Aws.Cryptography.StructuredEncryption.Types; namespace AWS.Cryptography.StructuredEncryption { - public class StructuredEncryption { - private readonly IStructuredEncryptionClient _impl; - public StructuredEncryption(AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig input) - { - Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredEncryptionConfig internalInput = TypeConversion.ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_StructuredEncryptionConfig(input); - var result = Dafny.Aws.Cryptography.StructuredEncryption.__default.StructuredEncryption(internalInput); - if (result.is_Failure) throw TypeConversion.FromDafny_CommonError(result.dtor_error); - this._impl = result.dtor_value; -} - public AWS.Cryptography.StructuredEncryption.EncryptStructureOutput EncryptStructure(AWS.Cryptography.StructuredEncryption.EncryptStructureInput input) { - Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput internalInput = TypeConversion.ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput(input); - Wrappers_Compile._IResult result = _impl.EncryptStructure(internalInput); - if (result.is_Failure) throw TypeConversion.FromDafny_CommonError(result.dtor_error); - return TypeConversion.FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput(result.dtor_value); -} - public AWS.Cryptography.StructuredEncryption.DecryptStructureOutput DecryptStructure(AWS.Cryptography.StructuredEncryption.DecryptStructureInput input) { - Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureInput internalInput = TypeConversion.ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput(input); - Wrappers_Compile._IResult result = _impl.DecryptStructure(internalInput); - if (result.is_Failure) throw TypeConversion.FromDafny_CommonError(result.dtor_error); - return TypeConversion.FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput(result.dtor_value); -} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionConfig.cs b/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionConfig.cs deleted file mode 100644 index c2185994b..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionConfig.cs +++ /dev/null @@ -1,13 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { - public class StructuredEncryptionConfig { - - - public void Validate() { - -} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionException.cs b/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionException.cs deleted file mode 100644 index fa6fb4897..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/StructuredEncryptionException.cs +++ /dev/null @@ -1,9 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System; - using AWS.Cryptography.StructuredEncryption; namespace AWS.Cryptography.StructuredEncryption { - public class StructuredEncryptionException : Exception { - public StructuredEncryptionException(string message) : base(message) {} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs b/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs deleted file mode 100644 index 86ccd8d13..000000000 --- a/src/StructuredEncryption/runtimes/net/Generated/TypeConversion.cs +++ /dev/null @@ -1,402 +0,0 @@ -// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. -// SPDX-License-Identifier: Apache-2.0 -// Do not modify this file. This file is machine generated, and any changes to it will be overwritten. -using System.Linq; namespace AWS.Cryptography.StructuredEncryption { - internal static class TypeConversion { - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M17_encryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { - return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); -} - public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M17_encryptionContext (System.Collections.Generic.Dictionary value) { - return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); -} - public static System.Collections.Generic.List FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M30_requiredContextFieldsOnDecrypt (Wrappers_Compile._IOption>> value) { - return value.is_None ? (System.Collections.Generic.List) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList(value.Extract()); -} - public static Wrappers_Compile._IOption>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M30_requiredContextFieldsOnDecrypt (System.Collections.Generic.List value) { - return value == null ? Wrappers_Compile.Option>>.create_None() : Wrappers_Compile.Option>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList((System.Collections.Generic.List) value)); -} - public static string FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes (Dafny.ISequence value) { - System.Text.UTF8Encoding utf8 = new System.Text.UTF8Encoding(false, true); -return utf8.GetString(value.Elements); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes (string value) { - System.Text.UTF8Encoding utf8 = new System.Text.UTF8Encoding(false, true); -return Dafny.Sequence.FromArray(utf8.GetBytes(value)); -} - public static AWS.Cryptography.StructuredEncryption.StructuredEncryptionException FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException (Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_StructuredEncryptionException value) { - return new AWS.Cryptography.StructuredEncryption.StructuredEncryptionException ( - FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message(value.message) - ) ; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_StructuredEncryptionException ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException (AWS.Cryptography.StructuredEncryption.StructuredEncryptionException value) { - - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_StructuredEncryptionException ( - ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message(value.Message) - ) ; -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes (Wrappers_Compile._IOption, Dafny.ISequence>> value) { - return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes(value.Extract()); -} - public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes (System.Collections.Generic.Dictionary value) { - return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes((System.Collections.Generic.Dictionary) value)); -} - public static string FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value (Dafny.ISequence value) { - return FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value (string value) { - return ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes (Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction> value) { - return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value(pair.Cdr)); -} - public static Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes (System.Collections.Generic.Dictionary value) { - return Dafny.Map, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>.FromCollection(value.Select(pair => - new Dafny.Pair, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value(pair.Value)) -)); -} - public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData)value; AWS.Cryptography.StructuredEncryption.StructuredData converted = new AWS.Cryptography.StructuredEncryption.StructuredData(); converted.Content = (AWS.Cryptography.StructuredEncryption.StructuredDataContent) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content(concrete.content); - if (concrete.attributes.is_Some) converted.Attributes = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes(concrete.attributes); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData (AWS.Cryptography.StructuredEncryption.StructuredData value) { - System.Collections.Generic.Dictionary var_attributes = value.IsSetAttributes() ? value.Attributes : (System.Collections.Generic.Dictionary) null; - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredData ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content(value.Content) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M10_attributes(var_attributes) ) ; -} - public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static AWS.Cryptography.StructuredEncryption.CryptoSchemaContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent value) { - -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent value) { - -} - public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList__M6_member (Dafny.ISequence value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S9_Utf8Bytes(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList__M6_member (string value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S9_Utf8Bytes(value); -} - public static AWS.Cryptography.StructuredEncryption.StructuredDataContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData__M7_content (AWS.Cryptography.StructuredEncryption.StructuredDataContent value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent(value); -} - public static AWS.Cryptography.StructuredEncryption.CryptoAction FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction value) { - if (value.is_ENCRYPT__AND__SIGN) return AWS.Cryptography.StructuredEncryption.CryptoAction.ENCRYPT_AND_SIGN; - if (value.is_SIGN__ONLY) return AWS.Cryptography.StructuredEncryption.CryptoAction.SIGN_ONLY; - if (value.is_IGNORE) return AWS.Cryptography.StructuredEncryption.CryptoAction.IGNORE; -throw new System.ArgumentException("Invalid AWS.Cryptography.StructuredEncryption.CryptoAction value"); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction (AWS.Cryptography.StructuredEncryption.CryptoAction value) { - if (AWS.Cryptography.StructuredEncryption.CryptoAction.ENCRYPT_AND_SIGN.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_ENCRYPT__AND__SIGN(); - if (AWS.Cryptography.StructuredEncryption.CryptoAction.SIGN_ONLY.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_SIGN__ONLY(); - if (AWS.Cryptography.StructuredEncryption.CryptoAction.IGNORE.Equals(value)) return Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoAction.create_IGNORE(); -throw new System.ArgumentException("Invalid AWS.Cryptography.StructuredEncryption.CryptoAction value"); -} - public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); -} - public static AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_StructuredEncryptionConfig (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredEncryptionConfig value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig)value; AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig converted = new AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig(); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredEncryptionConfig ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_StructuredEncryptionConfig (AWS.Cryptography.StructuredEncryption.StructuredEncryptionConfig value) { - - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.StructuredEncryptionConfig ( ) ; -} - public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M17_encryptionContext (Wrappers_Compile._IOption, Dafny.ISequence>> value) { - return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext(value.Extract()); -} - public static Wrappers_Compile._IOption, Dafny.ISequence>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M17_encryptionContext (System.Collections.Generic.Dictionary value) { - return value == null ? Wrappers_Compile.Option, Dafny.ISequence>>.create_None() : Wrappers_Compile.Option, Dafny.ISequence>>.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext((System.Collections.Generic.Dictionary) value)); -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key (Dafny.ISequence value) { - return FromDafny_N6_smithy__N3_api__S6_String(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M3_key (string value) { - return ToDafny_N6_smithy__N3_api__S6_String(value); -} - public static AWS.Cryptography.StructuredEncryption.EncryptStructureInput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput)value; AWS.Cryptography.StructuredEncryption.EncryptStructureInput converted = new AWS.Cryptography.StructuredEncryption.EncryptStructureInput(); converted.PlaintextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(concrete.plaintextStructure); - converted.CryptoSchema = (AWS.Cryptography.StructuredEncryption.CryptoSchema) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(concrete.cryptoSchema); - if (concrete.keyring.is_Some) converted.Keyring = (AWS.Cryptography.MaterialProviders.IKeyring) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring(concrete.keyring); - if (concrete.cmm.is_Some) converted.Cmm = (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm(concrete.cmm); - if (concrete.encryptionContext.is_Some) converted.EncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M17_encryptionContext(concrete.encryptionContext); - if (concrete.requiredContextFieldsOnDecrypt.is_Some) converted.RequiredContextFieldsOnDecrypt = (System.Collections.Generic.List) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M30_requiredContextFieldsOnDecrypt(concrete.requiredContextFieldsOnDecrypt); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureInput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput (AWS.Cryptography.StructuredEncryption.EncryptStructureInput value) { - AWS.Cryptography.MaterialProviders.IKeyring var_keyring = value.IsSetKeyring() ? value.Keyring : (AWS.Cryptography.MaterialProviders.IKeyring) null; - AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager var_cmm = value.IsSetCmm() ? value.Cmm : (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null; - System.Collections.Generic.Dictionary var_encryptionContext = value.IsSetEncryptionContext() ? value.EncryptionContext : (System.Collections.Generic.Dictionary) null; - System.Collections.Generic.List var_requiredContextFieldsOnDecrypt = value.IsSetRequiredContextFieldsOnDecrypt() ? value.RequiredContextFieldsOnDecrypt : (System.Collections.Generic.List) null; - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure(value.PlaintextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M12_cryptoSchema(value.CryptoSchema) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring(var_keyring) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm(var_cmm) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M17_encryptionContext(var_encryptionContext) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M30_requiredContextFieldsOnDecrypt(var_requiredContextFieldsOnDecrypt) ) ; -} - public static AWS.Cryptography.MaterialProviders.IKeyring FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring (Wrappers_Compile._IOption value) { - return value.is_None ? (AWS.Cryptography.MaterialProviders.IKeyring) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference(value.Extract()); -} - public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M7_keyring (AWS.Cryptography.MaterialProviders.IKeyring value) { - return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference((AWS.Cryptography.MaterialProviders.IKeyring) value)); -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion (Dafny.ISequence value) { - return new string(value.Elements); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion (string value) { - return Dafny.Sequence.FromString(value); -} - public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm (Wrappers_Compile._IOption value) { - return value.is_None ? (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference(value.Extract()); -} - public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager value) { - return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference((AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) value)); -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message (Dafny.ISequence value) { - return FromDafny_N6_smithy__N3_api__S6_String(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException__M7_message (string value) { - return ToDafny_N6_smithy__N3_api__S6_String(value); -} - public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema)value; AWS.Cryptography.StructuredEncryption.CryptoSchema converted = new AWS.Cryptography.StructuredEncryption.CryptoSchema(); converted.Content = (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content(concrete.content); - if (concrete.attributes.is_Some) converted.Attributes = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes(concrete.attributes); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { - System.Collections.Generic.Dictionary var_attributes = value.IsSetAttributes() ? value.Attributes : (System.Collections.Generic.Dictionary) null; - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.CryptoSchema ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content(value.Content) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes(var_attributes) ) ; -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes (Wrappers_Compile._IOption, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>> value) { - return value.is_None ? (System.Collections.Generic.Dictionary) null : FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes(value.Extract()); -} - public static Wrappers_Compile._IOption, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M10_attributes (System.Collections.Generic.Dictionary value) { - return value == null ? Wrappers_Compile.Option, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>>.create_None() : Wrappers_Compile.Option, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction>>.create_Some(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes((System.Collections.Generic.Dictionary) value)); -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key (Dafny.ISequence value) { - return FromDafny_N6_smithy__N3_api__S6_String(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key (string value) { - return ToDafny_N6_smithy__N3_api__S6_String(value); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext (Dafny.IMap, Dafny.ISequence> value) { - return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value(pair.Cdr)); -} - public static Dafny.IMap, Dafny.ISequence> ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext (System.Collections.Generic.Dictionary value) { - return Dafny.Map, Dafny.ISequence>.FromCollection(value.Select(pair => - new Dafny.Pair, Dafny.ISequence>(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M5_value(pair.Value)) -)); -} - public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference (Dafny.Aws.Cryptography.MaterialProviders.Types.ICryptographicMaterialsManager value) { - if (value is NativeWrapper_CryptographicMaterialsManager nativeWrapper) return nativeWrapper._impl; -return new CryptographicMaterialsManager(value); - -} - public static Dafny.Aws.Cryptography.MaterialProviders.Types.ICryptographicMaterialsManager ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager value) { - switch (value) - { - case CryptographicMaterialsManager valueWithImpl: - return valueWithImpl._impl; - case CryptographicMaterialsManagerBase nativeImpl: - return new NativeWrapper_CryptographicMaterialsManager(nativeImpl); - default: - throw new System.ArgumentException( - "Custom implementations of CryptographicMaterialsManager must extend CryptographicMaterialsManagerBase."); -} -} - public static AWS.Cryptography.StructuredEncryption.DecryptStructureOutput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureOutput value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureOutput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureOutput)value; AWS.Cryptography.StructuredEncryption.DecryptStructureOutput converted = new AWS.Cryptography.StructuredEncryption.DecryptStructureOutput(); converted.PlaintextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure(concrete.plaintextStructure); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureOutput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput (AWS.Cryptography.StructuredEncryption.DecryptStructureOutput value) { - - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureOutput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_DecryptStructureOutput__M18_plaintextStructure(value.PlaintextStructure) ) ; -} - public static AWS.Cryptography.StructuredEncryption.StructuredDataContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent value) { - -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredDataContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_StructuredDataContent (AWS.Cryptography.StructuredEncryption.StructuredDataContent value) { - -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas (Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema> value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas(value); -} - public static Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas (System.Collections.Generic.Dictionary value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas(value); -} - public static AWS.Cryptography.StructuredEncryption.CryptoSchemaContent FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchemaContent ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema__M7_content (AWS.Cryptography.StructuredEncryption.CryptoSchemaContent value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaContent(value); -} - public static string FromDafny_N6_smithy__N3_api__S6_String (Dafny.ISequence value) { - return new string(value.Elements); -} - public static Dafny.ISequence ToDafny_N6_smithy__N3_api__S6_String (string value) { - return Dafny.Sequence.FromString(value); -} - public static System.IO.MemoryStream FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal (Dafny.ISequence value) { - return new System.IO.MemoryStream(value.Elements); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal (System.IO.MemoryStream value) { - return Dafny.Sequence.FromArray(value.ToArray()); -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key (Dafny.ISequence value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key (string value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S19_CryptoSchemaVersion(value); -} - public static AWS.Cryptography.StructuredEncryption.CryptoSchema FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value (AWS.Cryptography.StructuredEncryption.CryptoSchema value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoSchema(value); -} - public static AWS.Cryptography.MaterialProviders.IKeyring FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference (Dafny.Aws.Cryptography.MaterialProviders.Types.IKeyring value) { - if (value is NativeWrapper_Keyring nativeWrapper) return nativeWrapper._impl; -return new Keyring(value); - -} - public static Dafny.Aws.Cryptography.MaterialProviders.Types.IKeyring ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference (AWS.Cryptography.MaterialProviders.IKeyring value) { - switch (value) - { - case Keyring valueWithImpl: - return valueWithImpl._impl; - case KeyringBase nativeImpl: - return new NativeWrapper_Keyring(nativeImpl); - default: - throw new System.ArgumentException( - "Custom implementations of Keyring must extend KeyringBase."); -} -} - public static AWS.Cryptography.StructuredEncryption.CryptoAction FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value (Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoAction ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_CryptoSchemaAttributes__M5_value (AWS.Cryptography.StructuredEncryption.CryptoAction value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S12_CryptoAction(value); -} - public static AWS.Cryptography.StructuredEncryption.EncryptStructureOutput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureOutput value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureOutput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureOutput)value; AWS.Cryptography.StructuredEncryption.EncryptStructureOutput converted = new AWS.Cryptography.StructuredEncryption.EncryptStructureOutput(); converted.CiphertextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure(concrete.ciphertextStructure); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IEncryptStructureOutput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput (AWS.Cryptography.StructuredEncryption.EncryptStructureOutput value) { - - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.EncryptStructureOutput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S22_EncryptStructureOutput__M19_ciphertextStructure(value.CiphertextStructure) ) ; -} - public static string FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key (Dafny.ISequence value) { - return FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S17_EncryptionContext__M3_key (string value) { - return ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S9_Utf8Bytes(value); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes (Dafny.IMap, Dafny.ISequence> value) { - return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value(pair.Cdr)); -} - public static Dafny.IMap, Dafny.ISequence> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes (System.Collections.Generic.Dictionary value) { - return Dafny.Map, Dafny.ISequence>.FromCollection(value.Select(pair => - new Dafny.Pair, Dafny.ISequence>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value(pair.Value)) -)); -} - public static AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm (Wrappers_Compile._IOption value) { - return value.is_None ? (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference(value.Extract()); -} - public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M3_cmm (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager value) { - return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S38_CryptographicMaterialsManagerReference((AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) value)); -} - public static System.Collections.Generic.List FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList (Dafny.ISequence> value) { - return new System.Collections.Generic.List(value.Elements.Select(FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList__M6_member)); -} - public static Dafny.ISequence> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList (System.Collections.Generic.List value) { - return Dafny.Sequence>.FromArray(value.Select(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S26_EncryptionContextFieldList__M6_member).ToArray()); -} - public static System.IO.MemoryStream FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value (Dafny.ISequence value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal(value); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S24_StructuredDataAttributes__M5_value (System.IO.MemoryStream value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S8_Terminal(value); -} - public static AWS.Cryptography.StructuredEncryption.DecryptStructureInput FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput (Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureInput value) { - Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput concrete = (Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput)value; AWS.Cryptography.StructuredEncryption.DecryptStructureInput converted = new AWS.Cryptography.StructuredEncryption.DecryptStructureInput(); converted.CiphertextStructure = (AWS.Cryptography.StructuredEncryption.StructuredData) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure(concrete.ciphertextStructure); - converted.CryptoSchemas = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas(concrete.cryptoSchemas); - if (concrete.keyring.is_Some) converted.Keyring = (AWS.Cryptography.MaterialProviders.IKeyring) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring(concrete.keyring); - if (concrete.cmm.is_Some) converted.Cmm = (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm(concrete.cmm); - if (concrete.encryptionContext.is_Some) converted.EncryptionContext = (System.Collections.Generic.Dictionary) FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M17_encryptionContext(concrete.encryptionContext); return converted; -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IDecryptStructureInput ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput (AWS.Cryptography.StructuredEncryption.DecryptStructureInput value) { - AWS.Cryptography.MaterialProviders.IKeyring var_keyring = value.IsSetKeyring() ? value.Keyring : (AWS.Cryptography.MaterialProviders.IKeyring) null; - AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager var_cmm = value.IsSetCmm() ? value.Cmm : (AWS.Cryptography.MaterialProviders.ICryptographicMaterialsManager) null; - System.Collections.Generic.Dictionary var_encryptionContext = value.IsSetEncryptionContext() ? value.EncryptionContext : (System.Collections.Generic.Dictionary) null; - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.DecryptStructureInput ( ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M19_ciphertextStructure(value.CiphertextStructure) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M13_cryptoSchemas(value.CryptoSchemas) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring(var_keyring) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M3_cmm(var_cmm) , ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M17_encryptionContext(var_encryptionContext) ) ; -} - public static string FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S9_Utf8Bytes (Dafny.ISequence value) { - System.Text.UTF8Encoding utf8 = new System.Text.UTF8Encoding(false, true); -return utf8.GetString(value.Elements); -} - public static Dafny.ISequence ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S9_Utf8Bytes (string value) { - System.Text.UTF8Encoding utf8 = new System.Text.UTF8Encoding(false, true); -return Dafny.Sequence.FromArray(utf8.GetBytes(value)); -} - public static AWS.Cryptography.MaterialProviders.IKeyring FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring (Wrappers_Compile._IOption value) { - return value.is_None ? (AWS.Cryptography.MaterialProviders.IKeyring) null : FromDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference(value.Extract()); -} - public static Wrappers_Compile._IOption ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_DecryptStructureInput__M7_keyring (AWS.Cryptography.MaterialProviders.IKeyring value) { - return value == null ? Wrappers_Compile.Option.create_None() : Wrappers_Compile.Option.create_Some(ToDafny_N3_aws__N12_cryptography__N17_materialProviders__S16_KeyringReference((AWS.Cryptography.MaterialProviders.IKeyring) value)); -} - public static System.Collections.Generic.Dictionary FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas (Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema> value) { - return value.ItemEnumerable.ToDictionary(pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key(pair.Car), pair => FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value(pair.Cdr)); -} - public static Dafny.IMap, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema> ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas (System.Collections.Generic.Dictionary value) { - return Dafny.Map, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema>.FromCollection(value.Select(pair => - new Dafny.Pair, Dafny.Aws.Cryptography.StructuredEncryption.Types._ICryptoSchema>(ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M3_key(pair.Key), ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S13_CryptoSchemas__M5_value(pair.Value)) -)); -} - public static AWS.Cryptography.StructuredEncryption.StructuredData FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure (Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData value) { - return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IStructuredData ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S21_EncryptStructureInput__M18_plaintextStructure (AWS.Cryptography.StructuredEncryption.StructuredData value) { - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S14_StructuredData(value); -} - public static System.Exception FromDafny_CommonError(Dafny.Aws.Cryptography.StructuredEncryption.Types._IError value) { - switch(value) - { - case Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_StructuredEncryptionException dafnyVal: -return FromDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException(dafnyVal); - case Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Opaque dafnyVal: - return new OpaqueError(dafnyVal.obj); - default: - // The switch MUST be complete for _IError, so `value` MUST NOT be an _IError. (How did you get here?) - return new OpaqueError(); -} -} - public static Dafny.Aws.Cryptography.StructuredEncryption.Types._IError ToDafny_CommonError(System.Exception value) { - switch (value) - { - case AWS.Cryptography.StructuredEncryption.StructuredEncryptionException exception: - return ToDafny_N3_aws__N12_cryptography__N20_structuredEncryption__S29_StructuredEncryptionException(exception); - // OpaqueError is redundant, but listed for completeness. - case OpaqueError exception: - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Opaque(exception); - case System.Exception exception: - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Opaque(exception); - default: - // The switch MUST be complete for System.Exception, so `value` MUST NOT be an System.Exception. (How did you get here?) - return new Dafny.Aws.Cryptography.StructuredEncryption.Types.Error_Opaque(value); -} -} -} -} diff --git a/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj b/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj index f0533f1d7..477acd11b 100644 --- a/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj +++ b/src/StructuredEncryption/runtimes/net/Test/AwsStructuredEncryptionTest.csproj @@ -53,10 +53,10 @@ - - - - + + + + diff --git a/src/StructuredEncryption/src/Index.dfy b/src/StructuredEncryption/src/Index.dfy index af843ca0f..7f324f8d6 100644 --- a/src/StructuredEncryption/src/Index.dfy +++ b/src/StructuredEncryption/src/Index.dfy @@ -1,8 +1,9 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" +include "../Model/AwsCryptographyStructuredEncryptionTypes.dfy" include "Operations/EncryptStructureOperation.dfy" include "Operations/DecryptStructureOperation.dfy" -// TODO including UInt in this file causes build issues. It complains of duplicate UInt modules... -include "../Model/AwsCryptographyStructuredEncryptionTypes.dfy" module {:extern "Dafny.Aws.StructuredEncryption.StructuredEncryptionClient"} StructuredEncryptionClient { import opened Wrappers diff --git a/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy b/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy index 90b6f79ca..ec9b4e6b0 100644 --- a/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy +++ b/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy @@ -1,3 +1,5 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 include "../../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" include "../../Model/AwsCryptographyStructuredEncryptionTypes.dfy" diff --git a/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy b/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy index 6cc7c59f5..1df86c240 100644 --- a/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy +++ b/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy @@ -1,3 +1,5 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 include "../../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" include "../../Model/AwsCryptographyStructuredEncryptionTypes.dfy" diff --git a/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj b/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj index 8f35e2cb5..91d42c5b2 100644 --- a/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj +++ b/src/StructuredEncryption/src/StructuredEncryptionDafny.csproj @@ -46,10 +46,10 @@ - - - - + + + + diff --git a/src/StructuredEncryption/test/HappyCaseTests.dfy b/src/StructuredEncryption/test/HappyCaseTests.dfy index 89bd140ac..0a8eb679b 100644 --- a/src/StructuredEncryption/test/HappyCaseTests.dfy +++ b/src/StructuredEncryption/test/HappyCaseTests.dfy @@ -1,3 +1,5 @@ +// Copyright Amazon.com Inc. or its affiliates. All Rights Reserved. +// SPDX-License-Identifier: Apache-2.0 include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" include "../src/Index.dfy" include "../Model/AwsCryptographyStructuredEncryptionTypes.dfy" From dfa21c1066a26af6d1d25815c4534beb034f685b Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Mon, 22 Aug 2022 16:48:56 -0700 Subject: [PATCH 10/13] Add newlines EOF --- src/StructuredEncryption/model/StructuredEncryption.smithy | 2 +- src/StructuredEncryption/src/Index.dfy | 2 +- .../src/Operations/DecryptStructureOperation.dfy | 2 +- .../src/Operations/EncryptStructureOperation.dfy | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/StructuredEncryption/model/StructuredEncryption.smithy b/src/StructuredEncryption/model/StructuredEncryption.smithy index 6798b0e67..3a18ace3f 100644 --- a/src/StructuredEncryption/model/StructuredEncryption.smithy +++ b/src/StructuredEncryption/model/StructuredEncryption.smithy @@ -172,4 +172,4 @@ string CryptoSchemaVersion structure StructuredEncryptionException { @required message: String, -} \ No newline at end of file +} diff --git a/src/StructuredEncryption/src/Index.dfy b/src/StructuredEncryption/src/Index.dfy index 7f324f8d6..81d3f9a8f 100644 --- a/src/StructuredEncryption/src/Index.dfy +++ b/src/StructuredEncryption/src/Index.dfy @@ -80,4 +80,4 @@ module {:extern "Dafny.Aws.StructuredEncryption.StructuredEncryptionClient"} Str History.DecryptStructure := History.DecryptStructure + [Types.DafnyCallEvent(input, output)]; } } -} \ No newline at end of file +} diff --git a/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy b/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy index ec9b4e6b0..6d0e1b055 100644 --- a/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy +++ b/src/StructuredEncryption/src/Operations/DecryptStructureOperation.dfy @@ -55,4 +55,4 @@ module DecryptStructureOperation { var decryptOutput := Types.DecryptStructureOutput(plaintextStructure := stubbedStructure); output := Success(decryptOutput); } -} \ No newline at end of file +} diff --git a/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy b/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy index 1df86c240..37ba4f4df 100644 --- a/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy +++ b/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy @@ -60,4 +60,4 @@ module EncryptStructureOperation { var encryptOutput := Types.EncryptStructureOutput(ciphertextStructure := stubbedStructure); output := Success(encryptOutput); } -} \ No newline at end of file +} From 06f798f0bedc6379ea390f5b79acfeb43cc120b9 Mon Sep 17 00:00:00 2001 From: lavaleri <49660121+lavaleri@users.noreply.github.com> Date: Tue, 23 Aug 2022 09:20:28 -0700 Subject: [PATCH 11/13] Apply suggestions from code review Co-authored-by: Tony Knapp <5892063+texastony@users.noreply.github.com> --- .../model/AwsCryptographyStructuredEncryptionTypes.dfy | 2 +- .../src/Operations/EncryptStructureOperation.dfy | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy index 606cd35c5..47811789e 100644 --- a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy +++ b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy @@ -3,7 +3,7 @@ // Do not modify this file. This file is machine generated, and any changes to it will be overwritten. // TODO had to manually update the below file. Polymorph needs to be fixed to not assume location of this file include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/StandardLibrary.dfy" -// TODO had to manually update the below file. Polymorph needs to be fixed to not assuer location of this file +// TODO had to manually update the below file. Polymorph needs to be fixed to not assume location of this file include "../../../private-aws-encryption-sdk-dafny-staging/src/Util/UTF8.dfy" include "../../../private-aws-encryption-sdk-dafny-staging/src/AwsCryptographicMaterialProviders/src/Index.dfy" module {:extern "Dafny.Aws.Cryptography.StructuredEncryption.Types" } AwsCryptographyStructuredEncryptionTypes diff --git a/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy b/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy index 37ba4f4df..130bbe71b 100644 --- a/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy +++ b/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy @@ -33,7 +33,7 @@ module EncryptStructureOperation { && input.encryptionContext.Some? && forall k :: k in input.requiredContextFieldsOnDecrypt.value ==> k in input.encryptionContext.value.Keys), Types.Error.StructuredEncryptionException( - message := "Invalid input: A Keyring or CMM MUST be supplied on input.")); + message := "Invalid input: Required Encryption Context fields are not in the input Encryption Context.")); // TODO: Currently stubbed out to return a hardcoded StructuredData var stubbedBytes := [0x21, 0x64, 0x6c, 0x72, 0x6f, 0x77, 0x20, 0x2c, 0x6f, 0x6c, 0x6c, 0x65, 0x68]; From 2e9b93c0362ed13041a1c745fa00bcd67035f2a2 Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Tue, 23 Aug 2022 12:25:05 -0700 Subject: [PATCH 12/13] Add check for unique items in required context fields --- src/StructuredEncryption/src/Index.dfy | 22 ++++++++++--------- .../Operations/EncryptStructureOperation.dfy | 15 +++++++++++-- 2 files changed, 25 insertions(+), 12 deletions(-) diff --git a/src/StructuredEncryption/src/Index.dfy b/src/StructuredEncryption/src/Index.dfy index 81d3f9a8f..318049e27 100644 --- a/src/StructuredEncryption/src/Index.dfy +++ b/src/StructuredEncryption/src/Index.dfy @@ -8,6 +8,7 @@ include "Operations/DecryptStructureOperation.dfy" module {:extern "Dafny.Aws.StructuredEncryption.StructuredEncryptionClient"} StructuredEncryptionClient { import opened Wrappers import opened StandardLibrary + import Seq import Types = AwsCryptographyStructuredEncryptionTypes import EncryptStructureOperation import DecryptStructureOperation @@ -32,22 +33,23 @@ module {:extern "Dafny.Aws.StructuredEncryption.StructuredEncryptionClient"} Str { var encryptionContextFields := if input.encryptionContext.Some? then input.encryptionContext.value.Keys else {}; var requiredFields := if input.requiredContextFieldsOnDecrypt.Some? then input.requiredContextFieldsOnDecrypt.value else []; - && !(forall k :: k in requiredFields ==> k in encryptionContextFields) ==> output.Failure? - && ( - || (input.cmm.Some? && input.keyring.Some?) - || (input.cmm.None? && input.keyring.None?) - ) ==> output.Failure? + && (!(forall k :: k in requiredFields ==> k in encryptionContextFields) ==> output.Failure?) + && (!Seq.HasNoDuplicates(requiredFields) ==> output.Failure?) + && (( + || (input.cmm.Some? && input.keyring.Some?) + || (input.cmm.None? && input.keyring.None?) + ) ==> output.Failure?) } predicate DecryptStructureEnsuresPublicly( input: Types.DecryptStructureInput, output: Result) { - && ( - || (input.cmm.Some? && input.keyring.Some?) - || (input.cmm.None? && input.keyring.None?) - ) ==> output.Failure? - && |input.cryptoSchemas| <= 0 ==> output.Failure? + && (( + || (input.cmm.Some? && input.keyring.Some?) + || (input.cmm.None? && input.keyring.None?) + ) ==> output.Failure?) + && (|input.cryptoSchemas| <= 0 ==> output.Failure?) } method EncryptStructure(input: Types.EncryptStructureInput) diff --git a/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy b/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy index 130bbe71b..8d36dda2d 100644 --- a/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy +++ b/src/StructuredEncryption/src/Operations/EncryptStructureOperation.dfy @@ -6,6 +6,7 @@ include "../../Model/AwsCryptographyStructuredEncryptionTypes.dfy" module EncryptStructureOperation { import opened Wrappers import opened StandardLibrary + import Seq import Types = AwsCryptographyStructuredEncryptionTypes method EncryptStructure(input: Types.EncryptStructureInput) @@ -18,6 +19,9 @@ module EncryptStructureOperation { var encryptionContextFields := if input.encryptionContext.Some? then input.encryptionContext.value.Keys else {}; var requiredFields := if input.requiredContextFieldsOnDecrypt.Some? then input.requiredContextFieldsOnDecrypt.value else []; !(forall k :: k in requiredFields ==> k in encryptionContextFields) ==> output.Failure? + ensures + var requiredFields := if input.requiredContextFieldsOnDecrypt.Some? then input.requiredContextFieldsOnDecrypt.value else []; + !Seq.HasNoDuplicates(requiredFields) ==> output.Failure? { // Ensure a Keyring XOR a CMM was included on input, otherwise error :- Need(input.cmm.Some? || input.keyring.Some?, @@ -27,14 +31,21 @@ module EncryptStructureOperation { Types.Error.StructuredEncryptionException( message := "Invalid input: Cannot recieve both a Keyring and a CMM on input.")); + var requiredContextFields := if input.requiredContextFieldsOnDecrypt.Some? then input.requiredContextFieldsOnDecrypt.value else []; // TODO Once verified, this information should probably move to some new EC datatype that combines this info // This datatype should live in the MPL. - :- Need(input.requiredContextFieldsOnDecrypt.Some? ==> ( + :- Need(|requiredContextFields| > 0 ==> ( && input.encryptionContext.Some? - && forall k :: k in input.requiredContextFieldsOnDecrypt.value ==> k in input.encryptionContext.value.Keys), + && forall k :: k in requiredContextFields ==> k in input.encryptionContext.value.Keys), Types.Error.StructuredEncryptionException( message := "Invalid input: Required Encryption Context fields are not in the input Encryption Context.")); + // TODO this check may move into the generated code if we build in @uniqueItems support + :- Need(|Seq.ToSet(requiredContextFields)| == |requiredContextFields|, + Types.Error.StructuredEncryptionException( + message := "Invalid input: Required Encryption Context fields contain duplicates.")); + Seq.LemmaNoDuplicatesCardinalityOfSet(requiredContextFields); + // TODO: Currently stubbed out to return a hardcoded StructuredData var stubbedBytes := [0x21, 0x64, 0x6c, 0x72, 0x6f, 0x77, 0x20, 0x2c, 0x6f, 0x6c, 0x6c, 0x65, 0x68]; var stubbedStructure := Types.StructuredData( From abdfb5c50821c811b916ccf41bf1eb3c08b1a9f6 Mon Sep 17 00:00:00 2001 From: Valerie Lambert Date: Tue, 23 Aug 2022 12:28:27 -0700 Subject: [PATCH 13/13] Add length trait to cryptoSchemas --- .../model/AwsCryptographyStructuredEncryptionTypes.dfy | 5 ++++- src/StructuredEncryption/model/StructuredEncryption.smithy | 6 ++++-- 2 files changed, 8 insertions(+), 3 deletions(-) diff --git a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy index 47811789e..8fbacf4f5 100644 --- a/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy +++ b/src/StructuredEncryption/model/AwsCryptographyStructuredEncryptionTypes.dfy @@ -33,7 +33,10 @@ include "../../../private-aws-encryption-sdk-dafny-staging/src/StandardLibrary/S | listSchema(CryptoSchemaList: CryptoSchemaList) type CryptoSchemaList = seq type CryptoSchemaMap = map - type CryptoSchemas = map + type CryptoSchemas = x: map | IsValid_CryptoSchemas(x) witness * + predicate method IsValid_CryptoSchemas(x: map) { + ( 1 <= |x| ) +} type CryptoSchemaVersion = string datatype DecryptStructureInput = | DecryptStructureInput ( nameonly ciphertextStructure: StructuredData , diff --git a/src/StructuredEncryption/model/StructuredEncryption.smithy b/src/StructuredEncryption/model/StructuredEncryption.smithy index 3a18ace3f..83ec639bf 100644 --- a/src/StructuredEncryption/model/StructuredEncryption.smithy +++ b/src/StructuredEncryption/model/StructuredEncryption.smithy @@ -46,7 +46,7 @@ structure EncryptStructureInput { cmm: CryptographicMaterialsManagerReference, encryptionContext: EncryptionContext, - contextFieldsRequiredOnDecrypt: EncryptionContextFieldList + requiredContextFieldsOnDecrypt: EncryptionContextFieldList } structure EncryptStructureOutput { @@ -73,7 +73,8 @@ structure DecryptStructureOutput { } // TODO move to MPL -// TODO this is better represented as a set +// TODO Until Polymorph supports this trait, verify uniqueness in Dafny +// @uniqueItems list EncryptionContextFieldList { member: Utf8Bytes } @@ -158,6 +159,7 @@ map CryptoSchemaAttributes { value: CryptoAction } +@length(min: 1) map CryptoSchemas { key: CryptoSchemaVersion, value: CryptoSchema