Skip to content

Commit

Permalink
Merge pull request #5 from Barkhausen-Institut/functional_correctness
Browse files Browse the repository at this point in the history
Functional correctness
  • Loading branch information
sertel authored May 7, 2024
2 parents af41708 + b573013 commit 44eae8c
Show file tree
Hide file tree
Showing 3 changed files with 109 additions and 29 deletions.
6 changes: 0 additions & 6 deletions theories/examples/RA.v
Original file line number Diff line number Diff line change
Expand Up @@ -568,8 +568,6 @@ Module Type RemoteAttestationHash
auto_in_fset.
Qed.

(*Print extructures.ord.tag_ordType.*)

Lemma INV'_full_heap_eq'_get : forall s1 s2,
full_heap_eq' (s1, s2) ->
∀ l,
Expand Down Expand Up @@ -728,15 +726,11 @@ Module Type RemoteAttestationHash
injective f -> (* if this is bijective then I would not end up in omap! *)
fmap_kmap' f (setm m k v) = setm (fmap_kmap' f m) (f k) v.
Proof.
Print fmap_kmap'.
move => f k v m inj_f.
rewrite /fmap_kmap'.
Fail rewrite [X in _ = setm _ X]mapm2E.
(* TODO *)
Check eq_fmap.
(* rewrite -eq_fmap. *)
Locate "=1".
Print eqfun.

(** * Approach 1:

Expand Down
100 changes: 92 additions & 8 deletions theories/examples/Sig_Prot.v
Original file line number Diff line number Diff line change
Expand Up @@ -65,25 +65,26 @@ Module Type SignatureProt

Import π1 π2 KG Alg Prim.

Definition Sig_prot_ifce :=
[interface #val #[sign] : 'message → 'pubkey × ('signature × 'bool) ].
Definition protocol := 30.
Definition Sig_prot_ifce :=
[interface #val #[protocol] : 'message → 'pubkey × ('signature × 'bool) ].

Definition Sig_prot : package Sig_locs_real Sig_ifce Sig_prot_ifce
:= [package
#def #[sign] (msg : 'message) : 'pubkey × ('signature × 'bool)
#def #[protocol] (msg : 'message) : 'pubkey × ('signature × 'bool)
{
#import {sig #[get_pk] : 'unit → 'pubkey } as get_pk ;;
#import {sig #[sign] : 'message → 'signature } as sign ;;
#import {sig #[verify_sig] : ('signature × 'message) → 'bool } as verify_sig ;;

(* Protocol *)
pk ← get_pk tt ;;
sig ← sign msg ;;
bool ← verify_sig (sig, msg) ;;
ret (pk, ( sig, bool ))
}
}
].


Equations Sig_prot_real : package Sig_locs_real [interface] Sig_prot_ifce :=
Sig_prot_real := {package Sig_prot ∘ Sig_real_c }.
Expand Down Expand Up @@ -129,7 +130,7 @@ Module Type SignatureProt
Defined.

Lemma ext_unforge_sig_prot:
Sig_prot_real ≈₀ Sig_prot_ideal.
Sig_prot_real ≈₀ Sig_prot_ideal.
Proof.
eapply (eq_rel_perf_ind_ignore (fset [:: sign_loc])).
- rewrite /Sig_locs_real/Sig_locs_ideal.
Expand Down Expand Up @@ -157,6 +158,89 @@ Module Type SignatureProt
apply Signature_prop.
---- by [move: pre; rewrite /inv_conj; repeat case].
Qed.


Module Correctness.

Definition prot_res := 100.

Definition Prot_res_ifce :=
[interface #val #[prot_res] : 'message → 'unit ].


Equations prot_result (msg : 'message) : code Sig_locs_real Sig_prot_ifce 'bool :=
prot_result msg := {code
#import {sig #[protocol] : 'message → 'pubkey × ('signature × 'bool) } as protocol ;;
'(pk, t) ← protocol msg;;
let '(_, result) := t in
ret result
}.

(* FIXME This just cannot simplify because it is not clear what the import is! *)
Theorem prot_correct seed msg:
Run sampler (prot_result msg) seed = Some true.
Proof.
simpl.
Admitted.


Equations prot_result_pkg : package Sig_locs_real Sig_prot_ifce Prot_res_ifce
:=
prot_result_pkg := [package
#def #[prot_res] (msg : 'message) : 'unit
{
#import {sig #[protocol] : 'message → 'pubkey × ('signature × 'bool) } as protocol ;;
'(_, t) ← protocol msg;;
let '(_, result) := t in
ret tt
}
].

(* TODO Why do I need this cast here? *)
Definition tt_ : chElement 'unit := tt.

Equations prot_result_pkg' : package Sig_locs_real Sig_prot_ifce Prot_res_ifce
:=
prot_result_pkg' := [package
#def #[prot_res] (msg : 'message) : 'unit
{
#import {sig #[protocol] : 'message → 'pubkey × ('signature × 'bool) } as protocol ;;
'(_, t) ← protocol msg;;
let '(_, result) := t in
#assert (result == true) ;;
ret tt_
}
].

Equations prot_result_real : package Sig_locs_real [interface] Prot_res_ifce :=
prot_result_real := {package prot_result_pkg ∘ Sig_prot ∘ Sig_real_c }.
Next Obligation.
ssprove_valid.
all: by [apply: fsubsetxx].
Defined.

Equations prot_result_real' : package Sig_locs_real [interface] Prot_res_ifce :=
prot_result_real' := {package prot_result_pkg' ∘ Sig_prot ∘ Sig_real_c }.
Next Obligation.
ssprove_valid.
all: by [apply: fsubsetxx].
Defined.

Lemma fun_correct:
prot_result_real ≈₀ prot_result_real'.
Proof.
eapply eq_rel_perf_ind_eq.
simplify_eq_rel x.
all: simplify_linking; ssprove_code_simpl.
repeat ssprove_sync_eq.
move => _.
ssprove_sync_eq => sk.
ssprove_sync_eq => pk.
rewrite /tt_.
rewrite (Signature_correct pk sk x) /=.
apply r_ret => s0 s1 s0_eq_s1 //=.
Qed.

End Correctness.

End SignatureProt.

32 changes: 17 additions & 15 deletions theories/examples/Signature.v
Original file line number Diff line number Diff line change
Expand Up @@ -75,8 +75,8 @@ End SignatureConstraints.
(** | KEY |
| GENERATION | **)

Module Type KeyGeneration
(π1 : SignatureParams)
Module Type KeyGeneration
(π1 : SignatureParams)
(π2 : SignatureConstraints).

Import π1 π2.
Expand Down Expand Up @@ -125,47 +125,50 @@ Module Type KeyGeneration
Definition Key_Gen : package Key_locs [interface] KeyGen_ifce
:= [package
#def #[key_gen] (_ : 'unit) : ('seckey × 'pubkey)
{
let (sk, pk) := KeyGen in
{
let (sk, pk) := KeyGen in
#put sk_loc := sk ;;
#put pk_loc := pk ;;
ret (sk, pk)

ret (sk, pk)
}
].

End KeyGeneration.

(** | SIGNATURE |
| SCHEME | **)

Module Type SignatureAlgorithms
(π1 : SignatureParams)
(π2 : SignatureConstraints)
Module Type SignatureAlgorithms
(π1 : SignatureParams)
(π2 : SignatureConstraints)
(π3 : KeyGeneration π1 π2).

Import π1 π2 π3.

Parameter Sign : ∀ (sk : SecKey) (m : chMessage), Signature.

Parameter Ver_sig : ∀ (pk : PubKey) (sig : Signature) (m : chMessage),
Parameter Ver_sig : ∀ (pk : PubKey) (sig : Signature) (m : chMessage),
'bool.

(* TODO: fmap (Signature * A * A ) -> (Signature * A * A ) triggert endless loop *)

(* Final proposition for a signature scheme to be indistinguishable *)
Parameter Signature_prop:
∀ (l: {fmap (Signature * chMessage ) -> 'unit})
∀ (l: {fmap (Signature * chMessage ) -> 'unit})
(s : Signature) (pk : PubKey) (m : chMessage),
Ver_sig pk s m = ((s,m) \in domm l).

(* Functional correctness property for signatures *)
Parameter Signature_correct: forall pk sk msg, Ver_sig pk (Sign sk msg) msg == true.

End SignatureAlgorithms.

Module Type SignaturePrimitives
(π1 : SignatureParams)
(π2 : SignatureConstraints)
(KG : KeyGeneration π1 π2)
(Alg : SignatureAlgorithms π1 π2 KG).
(Alg : SignatureAlgorithms π1 π2 KG).

Import π1 π2 KG Alg.

Expand All @@ -183,7 +186,7 @@ Module Type SignaturePrimitives

(* The signature scheme requires a heap location to store the seen signatures. *)
Definition Sig_locs_real := Key_locs.
Definition Sig_locs_ideal := Sig_locs_real :|: fset [:: sign_loc ].
Definition Sig_locs_ideal := Sig_locs_real :|: fset [:: sign_loc ].

Definition Sig_ifce := [interface
#val #[get_pk] : 'unit → 'pubkey ;
Expand Down Expand Up @@ -288,7 +291,6 @@ Module Type SignaturePrimitives
Sig_real_c ≈₀ Sig_ideal_c.
Proof.
eapply (eq_rel_perf_ind_ignore (fset [:: sign_loc])).
Check (_ :|: _).
- rewrite /Sig_locs_real/Sig_locs_ideal/Key_locs/Sig_locs_real/Key_locs.
apply fsubsetU.
apply/orP; right.
Expand Down

0 comments on commit 44eae8c

Please sign in to comment.