symm/ocb3.h, symm/ocb3-def.h: Implement the OCB3 auth'ned encryption mode.
[catacomb] / symm / ocb3.h
diff --git a/symm/ocb3.h b/symm/ocb3.h
new file mode 100644 (file)
index 0000000..ceb2df0
--- /dev/null
@@ -0,0 +1,337 @@
+/* -*-c-*-
+ *
+ * The OCB3 authenticated encryption mode
+ *
+ * (c) 2018 Straylight/Edgeware
+ */
+
+/*----- Licensing notice --------------------------------------------------*
+ *
+ * This file is part of Catacomb.
+ *
+ * Catacomb is free software: you can redistribute it and/or modify it
+ * under the terms of the GNU Library General Public License as published
+ * by the Free Software Foundation; either version 2 of the License, or
+ * (at your option) any later version.
+ *
+ * Catacomb is distributed in the hope that it will be useful, but
+ * WITHOUT ANY WARRANTY; without even the implied warranty of
+ * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the GNU
+ * Library General Public License for more details.
+ *
+ * You should have received a copy of the GNU Library General Public
+ * License along with Catacomb.  If not, write to the Free Software
+ * Foundation, Inc., 59 Temple Place - Suite 330, Boston, MA 02111-1307,
+ * USA.
+ */
+
+/*----- Notes on OCB3 -----------------------------------------------------*
+ *
+ * OCB3 was designed in 2011 by Phillip Rogaway and Ted Krovetz, building ten
+ * years of earlier work by Rogaway, Bellare, and Black, and refining Jutla's
+ * pioneering work on IACBC and IAPM.  OCB3 is an efficient `authenticated
+ * encryption with associated data' (AEAD) scheme which requires only one
+ * blockcipher application per message block.
+ *
+ * The patent situation on these efficient authenticated encryption schemes
+ * is fraught.  IBM hold two patents on Jutla's pioneering work on `IACBC'
+ * and `IAPM' which can apply (a third was filed at least six years too
+ * late), and Virgil Gligor and Pompiliu Donescu hold patents on their `XECB'
+ * and `XCBC' modes; these may or may not apply to OCB.  Rogaway himself
+ * holds US patents on various versions of OCB, but has issued free licences
+ * for free (`open source') software, and for all non-military use.  I think
+ * Catacomb's implementation of OCB falls within the scope of the former
+ * licence.
+ *
+ * OCB3 has optimized for short messages with `similar' nonces, where
+ * `similar' means `all but the low bits in the last byte are equal'; exactly
+ * how many bits depends on the block length in a rather complicated manner.
+ * This implementation supports this optimization through @pre_ocb3reinit@
+ * (which compares the new nonce against the old one to see whether it can
+ * make use of the similarity), and, more directly, through @pre_ocb3step@,
+ * which (effectively) advances the nonce as a big-endian counter.
+ *
+ * OCB3 was originally defined only for 128-bit blockciphers, and extending
+ * it to other sizes is nontrivial, but this has been done by Ted Krovetz in
+ * I-D draft-krovetz-ocb-wideblock-00 (following initial prompting from, err,
+ * me).
+ *
+ * OCB3 is a fairly well-behaved AEAD mode.  It doesn't require
+ * precommentment to the header or message lengths, but does require
+ * precommitment to the tag length.  It permits header data to be processed
+ * independently of any message.  It only accepts nonces the same size as the
+ * underlying blockcipher's block size, and it buffers up to a whole block's
+ * worth of data internally, which somewhat complicates streaming.
+ */
+
+#ifndef CATACOMB_OCB3_H
+#define CATACOMB_OCB3_H
+
+#ifdef __cplusplus
+  extern "C" {
+#endif
+
+/*----- Header files ------------------------------------------------------*/
+
+#include <stddef.h>
+
+#include <mLib/bits.h>
+#include <mLib/buf.h>
+
+#ifndef CATACOMB_GAEAD_H
+#  include "gaead.h"
+#endif
+
+#ifndef CATACOMB_OCB_H
+#  include "ocb.h"
+#endif
+
+/*----- Macros ------------------------------------------------------------*/
+
+/* --- @OCB3_DECL@ --- *
+ *
+ * Arguments:  @PRE@, @pre@ = prefixes for the underlying block cipher
+ *
+ * Use:                Creates declarations for OCB3 message-authentication mode.
+ */
+
+#define OCB3_NSZMAX(PRE) (PRE##_BLKSZ - (PRE##_BLKSZ <= 16 ? 1 : 2))
+
+#define OCB3_DECL(PRE, pre)                                            \
+                                                                       \
+typedef struct pre##_ocb3key {                                         \
+  pre##_ctx ctx;                       /* Underlying cipher context */ \
+  uint32 lstar[PRE##_BLKSZ/4];         /* Partial-block mask */        \
+  uint32 ldollar[PRE##_BLKSZ/4];       /* Checksum mask */             \
+  uint32 lmask[OCB_NCALC][PRE##_BLKSZ/4]; /* Precalculated masks */    \
+} pre##_ocb3key;                                                       \
+                                                                       \
+typedef struct pre##_ocb3aadctx {                                      \
+  pre##_ocb3key k;                     /* Processed key material */    \
+  uint32 o[PRE##_BLKSZ/4];             /* Current offset */            \
+  uint32 a[PRE##_BLKSZ/4];             /* Accumulator state */         \
+  octet b[PRE##_BLKSZ];                        /* Input buffer */              \
+  unsigned long i;                     /* Block counter */             \
+  unsigned off;                                /* Offset into buffered data */ \
+} pre##_ocb3aadctx;                                                    \
+                                                                       \
+typedef struct pre##_ocb3ctx {                                         \
+  pre##_ocb3key k;                     /* Processed key material */    \
+  unsigned nix, tsz;                   /* Nonce index and tag size */  \
+  uint32 nbase[PRE##_BLKSZ/2];         /* Current base nonce */        \
+  uint32 nstretch[PRE##_BLKSZ/2];      /* Stretched nonce material */  \
+  uint32 o[PRE##_BLKSZ/4];             /* Current offset */            \
+  uint32 a[PRE##_BLKSZ/4];             /* Accumulator state */         \
+  octet b[PRE##_BLKSZ];                        /* Input buffer */              \
+  unsigned long i;                     /* Block counter */             \
+  unsigned off;                                /* Offset into buffered data */ \
+} pre##_ocb3ctx;                                                       \
+                                                                       \
+extern const octet pre##_ocb3noncesz[], pre##_ocb3tagsz[];             \
+                                                                       \
+/* --- @pre_ocb3setkey@ --- *                                          \
+ *                                                                     \
+ * Arguments:  @pre_ocb3key *key@ = pointer to key block to fill in    \
+ *             @const void *k@ = pointer to key material               \
+ *             @size_t ksz@ = size of key material                     \
+ *                                                                     \
+ * Returns:    ---                                                     \
+ *                                                                     \
+ * Use:                Initializes an OCB3 key block.                          \
+ */                                                                    \
+                                                                       \
+extern void pre##_ocb3setkey(pre##_ocb3key */*key*/,                   \
+                            const void */*k*/, size_t /*ksz*/);        \
+                                                                       \
+/* --- @pre_ocb3aadinit@ --- *                                         \
+ *                                                                     \
+ * Arguments:  @pre_ocb3aadctx *aad@ = pointer to AAD context          \
+ *             @const pre_ocb3key *key@ = pointer to key block         \
+ *                                                                     \
+ * Returns:    ---                                                     \
+ *                                                                     \
+ * Use:                Initializes an OCB3 AAD (`additional authenticated      \
+ *             data') context associated with a given key.  AAD        \
+ *             contexts can be copied and/or reused, saving time if    \
+ *             the AAD for number of messages has a common prefix.     \
+ *                                                                     \
+ *             The @key@ doesn't need to be kept around, though        \
+ *             usually there'll at least be another copy in some OCB3  \
+ *             operation context because the AAD on its own isn't much \
+ *             good.                                                   \
+ */                                                                    \
+                                                                       \
+extern void pre##_ocb3aadinit(pre##_ocb3aadctx */*aad*/,               \
+                             const pre##_ocb3key */*key*/);            \
+                                                                       \
+/* --- @pre_ocb3aadhash@ --- *                                         \
+ *                                                                     \
+ * Arguments:  @pre_ocb3aadctx *aad@ = pointer to AAD context          \
+ *             @const void *p@ = pointer to AAD material               \
+ *             @size_t sz@ = length of AAD material                    \
+ *                                                                     \
+ * Returns:    ---                                                     \
+ *                                                                     \
+ * Use:                Feeds AAD into the context.                             \
+ */                                                                    \
+                                                                       \
+extern void pre##_ocb3aadhash(pre##_ocb3aadctx */*aad*/,               \
+                             const void */*p*/, size_t /*sz*/);        \
+                                                                       \
+/* --- @pre_ocb3init@ --- *                                            \
+ *                                                                     \
+ * Arguments:  @pre_ocb3ctx *ctx@ = pointer to OCB3 context            \
+ *             @const pre_ocb3key *key@ = pointer to key block         \
+ *             @const void *n@ = pointer to nonce                      \
+ *             @size_t nsz@ = size of nonce                            \
+ *             @size_t tsz@ = tag length                               \
+ *                                                                     \
+ * Returns:    Zero on success, @-1@ if the nonce or tag length is     \
+ *             bad.                                                    \
+ *                                                                     \
+ * Use:                Initialize an OCB3 operation context with a given key.  \
+ *                                                                     \
+ *             The original key needn't be kept around any more.       \
+ */                                                                    \
+                                                                       \
+extern int pre##_ocb3init(pre##_ocb3ctx */*ctx*/,                      \
+                         const pre##_ocb3key */*k*/,                   \
+                         const void */*n*/, size_t /*nsz*/,            \
+                         size_t /*tsz*/);                              \
+                                                                       \
+/* --- @pre_ocb3reinit@ --- *                                          \
+ *                                                                     \
+ * Arguments:  @pre_ocb3ctx *ctx@ = pointer to OCB3 context            \
+ *             @const void *n@ = pointer to nonce                      \
+ *             @size_t nsz@ = size of nonce                            \
+ *             @size_t tsz@ = tag length                               \
+ *                                                                     \
+ * Returns:    Zero on success, @-1@ if the nonce or tag length is     \
+ *             bad.                                                    \
+ *                                                                     \
+ * Use:                Reinitialize an OCB3 operation context, changing the    \
+ *             nonce and/or tag length.                                \
+ */                                                                    \
+                                                                       \
+extern int pre##_ocb3reinit(pre##_ocb3ctx */*ctx*/,                    \
+                           const void */*n*/, size_t /*nsz*/,          \
+                           size_t /*tsz*/);                            \
+                                                                       \
+/* --- @pre_ocb3step@ --- *                                            \
+ *                                                                     \
+ * Arguments:  @pre_ocb3ctx *ctx@ = pointer to OCB3 context            \
+ *                                                                     \
+ * Returns:    ---                                                     \
+ *                                                                     \
+ * Use:                Reinitialize an OCB3 operation context, stepping to     \
+ *             the `next' nonce along.                                 \
+ */                                                                    \
+                                                                       \
+extern void pre##_ocb3step(pre##_ocb3ctx */*ctx*/);                    \
+                                                                       \
+/* --- @pre_ocb3encrypt@ --- *                                         \
+ *                                                                     \
+ * Arguments:  @pre_ocb3ctx *ctx@ = pointer to OCB3 operation context  \
+ *             @const void *src@ = pointer to plaintext message chunk  \
+ *             @size_t sz@ = size of the plaintext                     \
+ *             @buf *dst@ = a buffer to write the ciphertext to        \
+ *                                                                     \
+ * Returns:    Zero on success; @-1@ on failure.                       \
+ *                                                                     \
+ * Use:                Encrypts a chunk of a plaintext message, writing a      \
+ *             chunk of ciphertext to the output buffer and updating   \
+ *             the operation state.                                    \
+ *                                                                     \
+ *             Note that OCB3 delays output if its input is not a      \
+ *             whole number of blocks.  This means that the output     \
+ *             might be smaller or larger the input by up to the block \
+ *             size.                                                   \
+ */                                                                    \
+                                                                       \
+extern int pre##_ocb3encrypt(pre##_ocb3ctx */*ctx*/,                   \
+                            const void */*src*/, size_t /*sz*/,        \
+                            buf */*dst*/);                             \
+                                                                       \
+/* --- @pre_ocb3decrypt@ --- *                                         \
+ *                                                                     \
+ * Arguments:  @pre_ocb3ctx *ctx@ = pointer to OCB3 operation context  \
+ *             @const void *src@ = pointer to ciphertext message chunk \
+ *             @size_t sz@ = size of the ciphertext                    \
+ *             @buf *dst@ = a buffer to write the plaintext to         \
+ *                                                                     \
+ * Returns:    Zero on success; @-1@ on failure.                       \
+ *                                                                     \
+ * Use:                Decrypts a chunk of a ciphertext message, writing a     \
+ *             chunk of plaintext to the output buffer and updating    \
+ *             the operation state.                                    \
+ *                                                                     \
+ *             Note that OCB3 delays output if its input is not a      \
+ *             whole number of blocks.  This means that the output     \
+ *             might be smaller or larger the input by up to the block \
+ *             size.                                                   \
+ */                                                                    \
+                                                                       \
+extern int pre##_ocb3decrypt(pre##_ocb3ctx */*ctx*/,                   \
+                            const void */*src*/, size_t /*sz*/,        \
+                            buf */*dst*/);                             \
+                                                                       \
+/* --- @pre_ocb3encryptdone@ --- *                                     \
+ *                                                                     \
+ * Arguments:  @pre_ocb3ctx *ctx@ = pointer to an OCB3 context         \
+ *             @const pre_ocb3aadctx *aad@ = pointer to AAD context,   \
+ *                     or null                                         \
+ *             @buf *dst@ = buffer for remaining ciphertext            \
+ *             @void *tag@ = where to write the tag                    \
+ *             @size_t tsz@ = length of tag to store                   \
+ *                                                                     \
+ * Returns:    Zero on success; @-1@ on failure.                       \
+ *                                                                     \
+ * Use:                Completes an OCB3 encryption operation.  The @aad@      \
+ *             pointer may be null if there is no additional           \
+ *             authenticated data.  OCB3 delays output, so this will   \
+ *             cause any remaining buffered plaintext to be encrypted  \
+ *             and written to @dst@.  Anyway, the function will fail   \
+ *             if the output buffer is broken.                         \
+ */                                                                    \
+                                                                       \
+extern int pre##_ocb3encryptdone(pre##_ocb3ctx */*ctx*/,               \
+                                const pre##_ocb3aadctx */*aad*/,       \
+                                buf */*dst*/,                          \
+                                void */*tag*/, size_t /*tsz*/);        \
+                                                                       \
+/* --- @pre_ocb3decryptdone@ --- *                                     \
+ *                                                                     \
+ * Arguments:  @pre_ocb3ctx *ctx@ = pointer to an OCB3 context         \
+ *             @const pre_ocb3aadctx *aad@ = pointer to AAD context,   \
+ *                     or null                                         \
+ *             @buf *dst@ = buffer for remaining plaintext             \
+ *             @const void *tag@ = tag to verify                       \
+ *             @size_t tsz@ = length of tag                            \
+ *                                                                     \
+ * Returns:    @+1@ for complete success; @0@ if tag verification      \
+ *             failed; @-1@ for other kinds of errors.                 \
+ *                                                                     \
+ * Use:                Completes an OCB3 decryption operation.  The @aad@      \
+ *             pointer may be null if there is no additional           \
+ *             authenticated data.  OCB3 delays output, so this will   \
+ *             cause any remaining buffered ciphertext to be decrypted \
+ *             and written to @dst@.  Anyway, the function will fail   \
+ *             if the output buffer is broken.                         \
+ */                                                                    \
+                                                                       \
+extern int pre##_ocb3decryptdone(pre##_ocb3ctx */*ctx*/,               \
+                                const pre##_ocb3aadctx */*aad*/,       \
+                                buf */*dst*/,                          \
+                                const void */*tag*/, size_t /*tsz*/);  \
+                                                                       \
+/* --- Generic AEAD interface --- */                                   \
+                                                                       \
+extern const gcaead pre##_ocb3;
+
+/*----- That's all, folks -------------------------------------------------*/
+
+#ifdef __cplusplus
+  }
+#endif
+
+#endif