Skip to content

Commit

Permalink
Add some tests
Browse files Browse the repository at this point in the history
  • Loading branch information
Nadrieril committed May 28, 2024
1 parent f372e08 commit 22c59f2
Show file tree
Hide file tree
Showing 9 changed files with 203 additions and 0 deletions.
54 changes: 54 additions & 0 deletions tests/coq/misc/Issue194RecursiveStructProjector.v
Original file line number Diff line number Diff line change
@@ -0,0 +1,54 @@
(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *)
(** [issue_194_recursive_struct_projector] *)
Require Import Primitives.
Import Primitives.
Require Import Coq.ZArith.ZArith.
Require Import List.
Import ListNotations.
Local Open Scope Primitives_scope.
Module Issue194RecursiveStructProjector.

(** [issue_194_recursive_struct_projector::AVLNode]
Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 2:0-2:17 *)
Inductive AVLNode_t (T : Type) :=
| mkAVLNode_t :
T ->
option (AVLNode_t T) ->
option (AVLNode_t T) ->
AVLNode_t T
.

Arguments mkAVLNode_t { _ }.

Definition aVLNode_value {T : Type} (x : AVLNode_t T) :=
match x with | mkAVLNode_t x1 _ _ => x1 end
.

Notation "x2 .(aVLNode_value)" := (aVLNode_value x2) (at level 9).

Definition aVLNode_left {T : Type} (x : AVLNode_t T) :=
match x with | mkAVLNode_t _ x1 _ => x1 end
.

Notation "x2 .(aVLNode_left)" := (aVLNode_left x2) (at level 9).

Definition aVLNode_right {T : Type} (x : AVLNode_t T) :=
match x with | mkAVLNode_t _ _ x1 => x1 end
.

Notation "x2 .(aVLNode_right)" := (aVLNode_right x2) (at level 9).

(** [issue_194_recursive_struct_projector::get_val]:
Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 10:0-10:33 *)
Definition get_val (T : Type) (x : AVLNode_t T) : result T :=
Ok x.(aVLNode_value)
.

(** [issue_194_recursive_struct_projector::get_left]:
Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 14:0-14:43 *)
Definition get_left
(T : Type) (x : AVLNode_t T) : result (option (AVLNode_t T)) :=
Ok x.(aVLNode_left)
.

End Issue194RecursiveStructProjector.
24 changes: 24 additions & 0 deletions tests/fstar/misc/Issue194RecursiveStructProjector.fst
Original file line number Diff line number Diff line change
@@ -0,0 +1,24 @@
(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *)
(** [issue_194_recursive_struct_projector] *)
module Issue194RecursiveStructProjector
open Primitives

#set-options "--z3rlimit 50 --fuel 1 --ifuel 1"

(** [issue_194_recursive_struct_projector::AVLNode]
Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 2:0-2:17 *)
type aVLNode_t (t : Type0) =
{
value : t; left : option (aVLNode_t t); right : option (aVLNode_t t);
}

(** [issue_194_recursive_struct_projector::get_val]:
Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 10:0-10:33 *)
let get_val (t : Type0) (x : aVLNode_t t) : result t =
Ok x.value

(** [issue_194_recursive_struct_projector::get_left]:
Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 14:0-14:43 *)
let get_left (t : Type0) (x : aVLNode_t t) : result (option (aVLNode_t t)) =
Ok x.left

12 changes: 12 additions & 0 deletions tests/fstar/misc/Matches.fst
Original file line number Diff line number Diff line change
@@ -0,0 +1,12 @@
(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *)
(** [matches] *)
module Matches
open Primitives

#set-options "--z3rlimit 50 --fuel 1 --ifuel 1"

(** [matches::match_u32]:
Source: 'tests/src/matches.rs', lines 4:0-4:27 *)
let match_u32 (x : u32) : result u32 =
begin match x with | 0 -> Ok 0 | 1 -> Ok 1 | _ -> Ok 2 end

25 changes: 25 additions & 0 deletions tests/lean/InfiniteLoop.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,25 @@
-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS
-- [infinite_loop]
import Base
open Primitives

namespace infinite_loop

/- [infinite_loop::bar]:
Source: 'tests/src/infinite-loop.rs', lines 4:0-4:8 -/
def bar : Result Unit :=
Result.ok ()

/- [infinite_loop::foo]: loop 0:
Source: 'tests/src/infinite-loop.rs', lines 6:0-10:1 -/
divergent def foo_loop : Unit :=
do
let _ ← bar
foo_loop

/- [infinite_loop::foo]:
Source: 'tests/src/infinite-loop.rs', lines 6:0-6:8 -/
def foo : Result Unit :=
foo_loop

end infinite_loop
35 changes: 35 additions & 0 deletions tests/lean/Issue194RecursiveStructProjector.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,35 @@
-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS
-- [issue_194_recursive_struct_projector]
import Base
open Primitives

namespace issue_194_recursive_struct_projector

/- [issue_194_recursive_struct_projector::AVLNode]
Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 2:0-2:17 -/
inductive AVLNode (T : Type) :=
| mk : T → Option (AVLNode T) → Option (AVLNode T) → AVLNode T

@[simp, reducible]
def AVLNode.value {T : Type} (x : AVLNode T) :=
match x with | AVLNode.mk x1 _ _ => x1

@[simp, reducible]
def AVLNode.left {T : Type} (x : AVLNode T) :=
match x with | AVLNode.mk _ x1 _ => x1

@[simp, reducible]
def AVLNode.right {T : Type} (x : AVLNode T) :=
match x with | AVLNode.mk _ _ x1 => x1

/- [issue_194_recursive_struct_projector::get_val]:
Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 10:0-10:33 -/
def get_val (T : Type) (x : AVLNode T) : Result T :=
Result.ok x.value

/- [issue_194_recursive_struct_projector::get_left]:
Source: 'tests/src/issue-194-recursive-struct-projector.rs', lines 14:0-14:43 -/
def get_left (T : Type) (x : AVLNode T) : Result (Option (AVLNode T)) :=
Result.ok x.left

end issue_194_recursive_struct_projector
16 changes: 16 additions & 0 deletions tests/lean/Matches.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,16 @@
-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS
-- [matches]
import Base
open Primitives

namespace matches

/- [matches::match_u32]:
Source: 'tests/src/matches.rs', lines 4:0-4:27 -/
def match_u32 (x : U32) : Result U32 :=
match x with
| 0#u32 => Result.ok 0#u32
| 1#u32 => Result.ok 1#u32
| _ => Result.ok 2#u32

end matches
11 changes: 11 additions & 0 deletions tests/src/infinite-loop.rs
Original file line number Diff line number Diff line change
@@ -0,0 +1,11 @@
//@ [coq,fstar,lean] skip
//@ [coq,fstar] aeneas-args=-use-fuel
//@ [coq,fstar] subdir=misc
// FIXME: make it work
fn bar() {}

fn foo() {
loop {
bar()
}
}
16 changes: 16 additions & 0 deletions tests/src/issue-194-recursive-struct-projector.rs
Original file line number Diff line number Diff line change
@@ -0,0 +1,16 @@
//@ [coq,fstar] subdir=misc
struct AVLNode<T> {
value: T,
left: AVLTree<T>,
right: AVLTree<T>,
}

type AVLTree<T> = Option<Box<AVLNode<T>>>;

fn get_val<T>(x: AVLNode<T>) -> T {
x.value
}

fn get_left<T>(x: AVLNode<T>) -> AVLTree<T> {
x.left
}
10 changes: 10 additions & 0 deletions tests/src/matches.rs
Original file line number Diff line number Diff line change
@@ -0,0 +1,10 @@
//@ [coq] skip
//@ [coq,fstar] subdir=misc
//^ note: coq gives "invalid notation for pattern"
fn match_u32(x: u32) -> u32 {
match x {
0 => 0,
1 => 1,
_ => 2,
}
}

0 comments on commit 22c59f2

Please sign in to comment.