diff --git a/Arm/BitVec.lean b/Arm/BitVec.lean index 1c82f194..1e59e45a 100644 --- a/Arm/BitVec.lean +++ b/Arm/BitVec.lean @@ -335,6 +335,12 @@ def split (x : BitVec n) (e : Nat) (h : 0 < e): List (BitVec e) := example : split 0xabcd1234#32 8 (by omega) = [0xab#8, 0xcd#8, 0x12#8, 0x34#8] := by rfl +/-- Reverse a list of bit vectors and flatten the list. -/ +def revflat (x : List (BitVec n)) : BitVec (n * x.length) := + have h : n * x.reverse.length = n * x.length := by + simp only [List.length_reverse] + BitVec.cast h $ BitVec.flatten (List.reverse x) + /-- Get the width of a bitvector. -/ protected def width (_ : BitVec n) : Nat := n diff --git a/Proofs/AES-GCM/AESGCMEncKernelPre.lean b/Proofs/AES-GCM/AESGCMEncKernelPre.lean new file mode 100644 index 00000000..e8dc1687 --- /dev/null +++ b/Proofs/AES-GCM/AESGCMEncKernelPre.lean @@ -0,0 +1,14 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author(s): Yan Peng +-/ +import Tests.«AES-GCM».AESGCMEncKernelProgram +import Tactics.StepThms + +namespace AESGCMEncKernelProgram + +set_option maxHeartbeats 2000000 in +#genStepEqTheorems aes_gcm_enc_kernel_program + +end AESGCMEncKernelProgram diff --git a/Proofs/AES-GCM/AESGCMEncKernelSym.lean b/Proofs/AES-GCM/AESGCMEncKernelSym.lean new file mode 100644 index 00000000..23029cc4 --- /dev/null +++ b/Proofs/AES-GCM/AESGCMEncKernelSym.lean @@ -0,0 +1,54 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author(s): Yan Peng +-/ +import Arm.Memory.Separate +import Proofs.«AES-GCM».AESGCMEncKernelPre +import Tactics.Sym + +namespace AESGCMEncKernelProgram + +abbrev in_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 0#5) s +abbrev in_bits_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 1#5) s +abbrev Xi_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 3#5) s +abbrev ivec_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 4#5) s +abbrev key_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 5#5) s +abbrev rounds_addr (s : ArmState) : BitVec 64 := (r (StateField.GPR 5#5) s) + 240#64 +abbrev Htable_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 6#5) s +abbrev out_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 2#5) s + +set_option maxRecDepth 1000000 in +theorem aes_gcm_enc_kernel_program_run_xx (s0 sf : ArmState) + (h_s0_program : s0.program = aes_gcm_enc_kernel_program) + (h_s0_err : read_err s0 = .None) + (h_s0_pc : read_pc s0 = aes_gcm_enc_kernel_program.min) + (h_s0_sp_aligned : CheckSPAlignment s0) + (h_run : sf = run 139 s0) + -- AES256 rounds = 14, the address of rounds is stored in x5+240 + (h_rounds : 14#32 = read_mem_bytes 4 (rounds_addr s0) s0) + -- memory separation + (h_mem : Memory.Region.pairwiseSeparate + [⟨(in_addr s0), 128⟩, ⟨(in_bits_addr s0), 64⟩, + ⟨(Xi_addr s0), 128⟩, ⟨(ivec_addr s0), 128⟩, + ⟨(key_addr s0), 1984⟩, -- 240*8 bits key schedule and 64 bits rounds + ⟨(Htable_addr s0), 2048⟩, -- 16*128 bits, only first 12 elements are used + ⟨(out_addr s0), 128⟩ ]) + : read_err sf = .None := by + simp (config := {ground := true}) only at h_s0_pc + sym_n 138 + case h_s1_sp_aligned => + apply Aligned_BitVecAdd_64_4 + · assumption + · simp (config := {ground := true}) only [] + init_next_step h_run h_step_139 s139 + replace h_step_139 := h_step_139.symm + rw [aes_gcm_enc_kernel_program.stepi_eq_0x7cf838] at h_step_139<;> try assumption + -- Instruction at pc 0x7cf79c#64 : cmp x17, #0xc, compares rounds + -- against 12 and branch out to AES-128 if rounds is less than 12 (equals 10) + -- We need accumulated (non)effects from this instruction to get branch condition + -- It would be tedious to list all the (non)effects out in a simp tactic + -- simp [state_simp_rules, bitvec_rules, minimal_theory] at h_step_139 + sorry + +end AESGCMEncKernelProgram diff --git a/Proofs/AES-GCM/AESHWEncryptSym.lean b/Proofs/AES-GCM/AESHWEncryptSym.lean new file mode 100644 index 00000000..38f54838 --- /dev/null +++ b/Proofs/AES-GCM/AESHWEncryptSym.lean @@ -0,0 +1,169 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author(s): Yan Peng +-/ +import Tests.«AES-GCM».AESHWEncryptProgram +import Tactics.Sym +import Tactics.StepThms + +namespace AESHWEncryptProgram + +#genStepEqTheorems aes_hw_encrypt_program + +abbrev in_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 0#5) s +abbrev key_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 2#5) s +abbrev round_addr (s : ArmState) : BitVec 64 := (r (StateField.GPR 2#5) s) + 240#64 +abbrev out_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 1#5) s + +theorem aes_hw_encrypt_program_run_60 (s0 sf : ArmState) + (h_s0_program : s0.program = aes_hw_encrypt_program) + (h_s0_err : read_err s0 = .None) + (h_s0_pc : read_pc s0 = aes_hw_encrypt_program.min) + (h_s0_sp_aligned : CheckSPAlignment s0) + (h_run : sf = run 60 s0) + -- AES256 rounds = 14, the address of rounds is stored in x2 + (h_rounds : 14 = read_mem_bytes 4 (round_addr s0) s0) + -- memory separation + (h_mem : Memory.Region.pairwiseSeparate + [⟨(in_addr s0), 128⟩, + ⟨(key_addr s0), 1984⟩, -- 240*8 bits key schedule and 64 bits rounds + ⟨(out_addr s0), 128⟩ ]) + : read_err sf = .None := by + simp (config := {ground := true}) only at h_s0_pc + -- ^^ Still needed, because `aes_hw_encrypt_program.min` is somehow + -- unable to be reflected + -- TODO: branching condition currently not supported so manually unroll loops + -- sym_n 13 fails with "tactic 'rewrite' failed, motive is not type correct" + sym_n 12 + init_next_step h_run h_step_13 s13 + replace h_step_13 := h_step_13.symm + rw [aes_hw_encrypt_program.stepi_eq_0x79f5d0] at h_step_13<;> try assumption + simp only [h_s1_x3, h_s2_non_effects, h_s3_non_effects, h_s4_x3, + h_s5_non_effects, h_s6_non_effects, h_s7_non_effects, + h_s8_non_effects, h_s9_flagN, h_s9_flagV, h_s9_flagZ, + h_s9_flagC, h_s10_non_effects, h_s11_non_effects, + h_s12_non_effects, state_simp_rules, bitvec_rules, + minimal_theory] at h_step_13 + simp only [round_addr] at h_rounds + simp (config := {ground := true}) only [h_rounds.symm, minimal_theory] at h_step_13 + (intro_fetch_decode_lemmas h_step_13 h_s12_program "h_s12") + -- Add hypotheses that are needed for next loop iteration + -- This is an aggregated result + have h_s13_x3 : read_gpr 32 3#5 s13 = 10#32 := by + simp only [h_s1_x3, h_s2_non_effects, h_s3_non_effects, h_s4_x3, + h_s5_non_effects, h_s6_non_effects, h_s7_non_effects, + h_s8_non_effects, h_s9_x3, h_s10_non_effects, + h_s11_non_effects, h_s12_non_effects, h_step_13, + state_simp_rules, bitvec_rules, minimal_theory] + simp (config := {ground := true}) only [h_rounds.symm, minimal_theory] + -- + sym_n 7 at s13 + init_next_step h_run h_step_21 s21 + replace h_step_21 := h_step_21.symm + rw [aes_hw_encrypt_program.stepi_eq_0x79f5d0] at h_step_21<;> try assumption + simp only [state_simp_rules] at h_s13_x3 + simp only [h_s20_non_effects, h_s19_non_effects, h_s18_non_effects, + h_s17_flagN, h_s17_flagV, h_s17_flagZ, + h_s17_flagC, h_s16_non_effects, h_s15_non_effects, + h_s14_non_effects, state_simp_rules, bitvec_rules, + minimal_theory, + -- hypothesis that states x3's value + h_s13_x3] at h_step_21 + simp (config := {ground := true}) only [minimal_theory] at h_step_21 + (intro_fetch_decode_lemmas h_step_21 h_s20_program "h_s20") + -- Add hypotheses that are needed for next loop iteration + -- This is an aggregated result + have h_s21_x3 : read_gpr 32 3#5 s21 = 8#32 := by + simp only [h_s20_non_effects, h_s19_non_effects, h_s18_non_effects, + h_s17_x3, h_s16_non_effects, h_s15_non_effects, + h_s14_non_effects, h_step_21, + state_simp_rules, bitvec_rules, minimal_theory] + simp (config := {ground := true}) only [h_s13_x3, minimal_theory] + -- + sym_n 7 at s21 + init_next_step h_run h_step_29 s29 + replace h_step_29 := h_step_29.symm + rw [aes_hw_encrypt_program.stepi_eq_0x79f5d0] at h_step_29<;> try assumption + simp only [state_simp_rules] at h_s21_x3 + simp only [h_s28_non_effects, h_s27_non_effects, h_s26_non_effects, + h_s25_flagN, h_s25_flagV, h_s25_flagZ, + h_s25_flagC, h_s24_non_effects, h_s23_non_effects, + h_s22_non_effects, state_simp_rules, bitvec_rules, + minimal_theory, + -- hypothesis that states x3's value + h_s21_x3] at h_step_29 + simp (config := {ground := true}) only [minimal_theory] at h_step_29 + (intro_fetch_decode_lemmas h_step_29 h_s28_program "h_s28") + -- Add hypotheses that are needed for next loop iteration + -- This is an aggregated result + have h_s29_x3 : read_gpr 32 3#5 s29 = 6#32 := by + simp only [h_s28_non_effects, h_s27_non_effects, h_s26_non_effects, + h_s25_x3, h_s24_non_effects, h_s23_non_effects, + h_s22_non_effects, h_step_29, + state_simp_rules, bitvec_rules, minimal_theory] + simp (config := {ground := true}) only [h_s21_x3, minimal_theory] + -- + sym_n 7 at s29 + init_next_step h_run h_step_37 s37 + replace h_step_37 := h_step_37.symm + rw [aes_hw_encrypt_program.stepi_eq_0x79f5d0] at h_step_37<;> try assumption + simp only [state_simp_rules] at h_s29_x3 + simp only [h_s36_non_effects, h_s35_non_effects, h_s34_non_effects, + h_s33_flagN, h_s33_flagV, h_s33_flagZ, + h_s33_flagC, h_s32_non_effects, h_s31_non_effects, + h_s30_non_effects, state_simp_rules, bitvec_rules, + minimal_theory, + -- hypothesis that states x3's value + h_s29_x3] at h_step_37 + simp (config := {ground := true}) only [minimal_theory] at h_step_37 + (intro_fetch_decode_lemmas h_step_37 h_s36_program "h_s36") + -- Add hypotheses that are needed for next loop iteration + -- This is an aggregated result + have h_s37_x3 : read_gpr 32 3#5 s37 = 4#32 := by + simp only [h_s36_non_effects, h_s35_non_effects, h_s34_non_effects, + h_s33_x3, h_s32_non_effects, h_s31_non_effects, + h_s30_non_effects, h_step_37, + state_simp_rules, bitvec_rules, minimal_theory] + simp (config := {ground := true}) only [h_s29_x3, minimal_theory] + -- + sym_n 7 at s37 + init_next_step h_run h_step_45 s45 + replace h_step_45 := h_step_45.symm + rw [aes_hw_encrypt_program.stepi_eq_0x79f5d0] at h_step_45<;> try assumption + simp only [state_simp_rules] at h_s37_x3 + simp only [h_s44_non_effects, h_s43_non_effects, h_s42_non_effects, + h_s41_flagN, h_s41_flagV, h_s41_flagZ, + h_s41_flagC, h_s40_non_effects, h_s39_non_effects, + h_s38_non_effects, state_simp_rules, bitvec_rules, + minimal_theory, + -- hypothesis that states x3's value + h_s37_x3] at h_step_45 + simp (config := {ground := true}) only [minimal_theory] at h_step_45 + (intro_fetch_decode_lemmas h_step_45 h_s44_program "h_s44") + -- Add hypotheses that are needed for next loop iteration + -- This is an aggregated result + have h_s45_x3 : read_gpr 32 3#5 s45 = 2#32 := by + simp only [h_s44_non_effects, h_s43_non_effects, h_s42_non_effects, + h_s41_x3, h_s40_non_effects, h_s39_non_effects, + h_s38_non_effects, h_step_45, + state_simp_rules, bitvec_rules, minimal_theory] + simp (config := {ground := true}) only [h_s37_x3, minimal_theory] + -- + sym_n 7 at s45 + init_next_step h_run h_step_53 s53 + replace h_step_53 := h_step_53.symm + rw [aes_hw_encrypt_program.stepi_eq_0x79f5d0] at h_step_53<;> try assumption + simp only [state_simp_rules] at h_s45_x3 + simp only [h_s52_non_effects, h_s51_non_effects, h_s50_non_effects, + h_s49_flagN, h_s49_flagV, h_s49_flagZ, + h_s49_flagC, h_s48_non_effects, h_s47_non_effects, + h_s46_non_effects, state_simp_rules, bitvec_rules, + minimal_theory, + -- hypothesis that states x3's value + h_s45_x3] at h_step_53 + simp (config := {ground := true}) only [minimal_theory] at h_step_53 + (intro_fetch_decode_lemmas h_step_53 h_s52_program "h_s52") + -- + sym_n 7 at s53 + done diff --git a/Proofs/AES-GCM/GCM.lean b/Proofs/AES-GCM/GCM.lean index 2fbac3c8..53374f41 100644 --- a/Proofs/AES-GCM/GCM.lean +++ b/Proofs/AES-GCM/GCM.lean @@ -1 +1,9 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author(s): Alex Keizer +-/ import Proofs.«AES-GCM».GCMGmultV8Sym +import Proofs.«AES-GCM».GCMInitV8Sym +import Proofs.«AES-GCM».AESHWEncryptSym +import Proofs.«AES-GCM».AESGCMEncKernelSym diff --git a/Proofs/AES-GCM/GCMGmultV8Sym.lean b/Proofs/AES-GCM/GCMGmultV8Sym.lean index c3e35a70..7114063f 100644 --- a/Proofs/AES-GCM/GCMGmultV8Sym.lean +++ b/Proofs/AES-GCM/GCMGmultV8Sym.lean @@ -1,3 +1,8 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author(s): Alex Keizer +-/ import Tests.«AES-GCM».GCMGmultV8Program import Tactics.Sym import Tactics.StepThms diff --git a/Proofs/AES-GCM/GCMInitV8Pre.lean b/Proofs/AES-GCM/GCMInitV8Pre.lean new file mode 100644 index 00000000..505c532b --- /dev/null +++ b/Proofs/AES-GCM/GCMInitV8Pre.lean @@ -0,0 +1,14 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author(s): Yan Peng +-/ +import Tests.«AES-GCM».GCMInitV8Program +import Tactics.StepThms + +namespace GCMInitV8Program + +set_option maxHeartbeats 1000000 in +#genStepEqTheorems gcm_init_v8_program + +end GCMInitV8Program diff --git a/Proofs/AES-GCM/GCMInitV8Sym.lean b/Proofs/AES-GCM/GCMInitV8Sym.lean new file mode 100644 index 00000000..a17a8e41 --- /dev/null +++ b/Proofs/AES-GCM/GCMInitV8Sym.lean @@ -0,0 +1,89 @@ +/- +Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author(s): Yan Peng +-/ +import Arm.BitVec +import Proofs.«AES-GCM».GCMInitV8Pre +import Tactics.Sym +import Tactics.Aggregate +import Specs.GCMV8 + +namespace GCMInitV8Program + +abbrev H_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 1#5) s +abbrev Htable_addr (s : ArmState) : BitVec 64 := r (StateField.GPR 0#5) s + +-- set_option debug.byAsSorry true in +set_option maxRecDepth 1000000 in +-- set_option profiler true in +theorem gcm_init_v8_program_run_152 (s0 sf : ArmState) + (h_s0_program : s0.program = gcm_init_v8_program) + (h_s0_err : read_err s0 = .None) + (h_s0_pc : read_pc s0 = gcm_init_v8_program.min) + (h_s0_sp_aligned : CheckSPAlignment s0) + (h_run : sf = run gcm_init_v8_program.length s0) + (h_mem : Memory.Region.pairwiseSeparate + [ ⟨(H_addr s0), 128⟩, + ⟨(Htable_addr s0), 2048⟩ ]) + : read_err sf = .None := by + simp (config := {ground := true}) only at h_s0_pc + -- ^^ Still needed, because `gcm_init_v8_program.min` is somehow + -- unable to be reflected + sym_n 152 + done + +set_option maxRecDepth 1000000 in +set_option maxHeartbeats 2000000 in +-- set_option profiler true in +theorem gcm_init_v8_program_correct (s0 sf : ArmState) + (h_s0_program : s0.program = gcm_init_v8_program) + (h_s0_err : read_err s0 = .None) + (h_s0_pc : read_pc s0 = gcm_init_v8_program.min) + (h_s0_sp_aligned : CheckSPAlignment s0) + (h_run : sf = run gcm_init_v8_program.length s0) + (h_mem : Memory.Region.pairwiseSeparate + [ ⟨(H_addr s0), 128⟩, + ⟨(Htable_addr s0), 2048⟩ ]) + : -- effects + -- no instruction decoding error + read_err sf = .None + -- program is unchanged + ∧ sf.program = gcm_init_v8_program + -- SP is still aligned + ∧ CheckSPAlignment sf + -- PC is updated + ∧ read_pc sf = read_gpr 64 30#5 s0 + -- Htable_addr ptr is moved to the start of the 10th element + ∧ Htable_addr sf = Htable_addr s0 + (9 * 16#64) + -- H_addr ptr stays the same + ∧ H_addr sf = H_addr s0 + -- v20 - v31 stores results of Htable + ∧ read_sfp 128 20#5 sf = (GCMV8.GCMInitV8 (read_mem_bytes 16 (H_addr s0) s0)).get! 0 + -- ∧ read_sfp 128 21#5 sf = (GCMV8.GCMInitV8 (read_mem_bytes 16 (H_addr s0) s0)).get! 1 + -- TODO: Commenting out memory related conjuncts since it seems + -- to make symbolic execution stuck + -- -- First 12 elements in Htable is correct + -- ∧ read_mem_bytes 192 (Htable_addr s0) sf + -- = revflat (GCMV8.GCMInitV8 (read_mem_bytes 16 (H_addr s0) s0)) + -- + -- non-effects + -- State values that shouldn't change do not change + -- TODO: figure out all registers that are used ... + -- ∧ (∀ (f : StateField), ¬ (f = StateField.PC) ∧ + -- ¬ (f = (StateField.GPR 29#5)) → + -- r f sf = r f s0) + -- -- Memory safety: memory location that should not change did + -- -- not change + -- -- The addr exclude output region Htable + -- ∧ (∀ (n : Nat) (addr : BitVec 64) (h: addr < (Htable_addr sf) ∨ addr >= (Htable_addr sf) + 128*12), + -- read_mem_bytes n addr sf = read_mem_bytes n addr s0) + := by + simp (config := {ground := true}) only at h_s0_pc + -- ^^ Still needed, because `gcm_init_v8_program.min` is somehow + -- unable to be reflected + sym_n 152 + simp only [Htable_addr, state_value] -- TODO: state_value is needed, why + apply And.intro + · bv_decide + · sorry diff --git a/Tests/AES-GCM/AESGCMDecKernelProgram.lean b/Tests/AES-GCM/AESGCMDecKernelProgram.lean index 8d9131d5..06bd110d 100644 --- a/Tests/AES-GCM/AESGCMDecKernelProgram.lean +++ b/Tests/AES-GCM/AESGCMDecKernelProgram.lean @@ -21,7 +21,7 @@ open BitVec Xi -- current hash/tag value, should match with input to aes_gcm_enc_kernel (x3) ivec -- counter, should match with input to aes_gcm_enc_kernel (x4) key -- AES key schedule and rounds (x5) - Htable -- powers of H precomputed up to H^12 (x6) + Htable -- powers of H precomputed up to H^8 (x6) output: out -- plaintext (x2) Xi -- current hash/tag value, decryption function will recalculate the hash based on the ciphertext, the top-level c function diff --git a/Tests/AES-GCM/AESGCMEncKernelProgram.lean b/Tests/AES-GCM/AESGCMEncKernelProgram.lean index 193af54e..8a9c201b 100644 --- a/Tests/AES-GCM/AESGCMEncKernelProgram.lean +++ b/Tests/AES-GCM/AESGCMEncKernelProgram.lean @@ -21,7 +21,7 @@ open BitVec Xi -- current hash/tag value (x3) ivec -- counter (x4) key -- AES key schedule and rounds (x5) - Htable -- powers of H precomputed up to H^12 (x6) + Htable -- powers of H precomputed up to H^8 (x6) output: out -- ciphertext (x2) Xi -- current hash/tag value (x3) diff --git a/Tests/AES-GCM/AESGCMProgramTests.lean b/Tests/AES-GCM/AESGCMProgramTests.lean index f8b91725..34538ccf 100644 --- a/Tests/AES-GCM/AESGCMProgramTests.lean +++ b/Tests/AES-GCM/AESGCMProgramTests.lean @@ -6,7 +6,6 @@ Author(s): Yan Peng import Arm.BitVec import Arm.Exec import Arm.Cfg.Cfg -import Tests.Common import Tests.«AES-GCM».AESGCMEncKernelProgram import Tests.«AES-GCM».AESGCMDecKernelProgram diff --git a/Tests/AES-GCM/AESV8ProgramTests.lean b/Tests/AES-GCM/AESV8ProgramTests.lean index 7201add6..4389d415 100644 --- a/Tests/AES-GCM/AESV8ProgramTests.lean +++ b/Tests/AES-GCM/AESV8ProgramTests.lean @@ -6,7 +6,6 @@ Author(s): Yan Peng import Arm.BitVec import Arm.Exec import Arm.Cfg.Cfg -import Tests.Common import Tests.«AES-GCM».AESHWSetEncryptKeyProgram import Tests.«AES-GCM».AESHWEncryptProgram import Tests.«AES-GCM».AESHWCtr32EncryptBlocksProgram diff --git a/Tests/AES-GCM/GCMInitV8Program.lean b/Tests/AES-GCM/GCMInitV8Program.lean index 70c51518..10c43638 100644 --- a/Tests/AES-GCM/GCMInitV8Program.lean +++ b/Tests/AES-GCM/GCMInitV8Program.lean @@ -14,7 +14,7 @@ open BitVec void gcm_init_v8(u128 Htable[16], const uint64_t H[2]); input: H - initial hash (x1) - output: Htable - powers of H precomputed up to H^12 (x0) + output: Htable - powers of H precomputed up to H^8 (x0) -/ def gcm_init_v8_program : Program := diff --git a/Tests/AES-GCM/GCMProgramTests.lean b/Tests/AES-GCM/GCMProgramTests.lean index 27e502f7..16bc76fa 100644 --- a/Tests/AES-GCM/GCMProgramTests.lean +++ b/Tests/AES-GCM/GCMProgramTests.lean @@ -6,7 +6,6 @@ Author(s): Yan Peng import Arm.BitVec import Arm.Exec import Arm.Cfg.Cfg -import Tests.Common import Tests.«AES-GCM».GCMGhashV8Program import Tests.«AES-GCM».GCMGmultV8Program import Tests.«AES-GCM».GCMInitV8Program diff --git a/Tests/Common.lean b/Tests/Common.lean deleted file mode 100644 index 4129e20f..00000000 --- a/Tests/Common.lean +++ /dev/null @@ -1,10 +0,0 @@ -/- -Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Author(s): Yan Peng --/ -import Arm.BitVec - -def revflat (x : List (BitVec n)) : BitVec (n * x.length) := - have h : n * x.reverse.length = n * x.length := by simp only [List.length_reverse] - BitVec.cast h $ BitVec.flatten (List.reverse x)