mirror of
https://github.com/supleed2/ELEC70056-HSV-CW1.git
synced 2024-12-22 21:55:48 +00:00
releasing dafny/isabelle coursework
This commit is contained in:
parent
d33337a7a6
commit
c844f23181
BIN
dafny/2022/dafny_exercises_2022.pdf
Normal file
BIN
dafny/2022/dafny_exercises_2022.pdf
Normal file
Binary file not shown.
120
dafny/2022/tasks_2022.dfy
Normal file
120
dafny/2022/tasks_2022.dfy
Normal file
|
@ -0,0 +1,120 @@
|
|||
// Dafny coursework tasks
|
||||
// Autumn term, 2022
|
||||
//
|
||||
// Authors: John Wickerson
|
||||
|
||||
method countsquares(n:nat) returns (result:nat)
|
||||
//ensures result == ...
|
||||
{
|
||||
var i := 0;
|
||||
result := 0;
|
||||
while i < n {
|
||||
i := i + 1;
|
||||
result := result + i * i;
|
||||
}
|
||||
}
|
||||
|
||||
method countsquares2(n:nat) returns (result:nat)
|
||||
// ensures result == ...
|
||||
{
|
||||
var i := n;
|
||||
result := 0;
|
||||
while i > 0 {
|
||||
result := result + i * i;
|
||||
i := i - 1;
|
||||
}
|
||||
}
|
||||
|
||||
predicate sorted(A:array<int>)
|
||||
reads A
|
||||
{
|
||||
forall m,n :: 0 <= m < n < A.Length ==> A[m] <= A[n]
|
||||
}
|
||||
|
||||
method binarysearch_between(A:array<int>, v:int, lo:int, hi:int) returns (result:bool)
|
||||
requires false // Please remove this line.
|
||||
{
|
||||
if lo == hi {
|
||||
return false;
|
||||
}
|
||||
var mid:int := (lo + hi) / 2;
|
||||
if v == A[mid] {
|
||||
return true;
|
||||
}
|
||||
if v < A[mid] {
|
||||
result := binarysearch_between(A, v, lo, mid);
|
||||
}
|
||||
if v > A[mid] {
|
||||
result := binarysearch_between(A, v, mid+1, hi);
|
||||
}
|
||||
}
|
||||
|
||||
method binarysearch(A:array<int>, v:int) returns (result:bool)
|
||||
requires false // Please remove this line.
|
||||
{
|
||||
result := binarysearch_between(A, v, 0, A.Length);
|
||||
}
|
||||
|
||||
method binarysearch_iter(A:array<int>, v:int) returns (result:bool)
|
||||
{
|
||||
// Please provide an implementation of this method.
|
||||
}
|
||||
|
||||
method partition(A:array<int>, lo:int, hi:int) returns (pivot:int)
|
||||
requires 0 <= lo < hi <= A.Length
|
||||
ensures 0 <= lo <= pivot < hi
|
||||
ensures forall k :: (0 <= k < lo || hi <= k < A.Length) ==> old(A[k]) == A[k]
|
||||
modifies A
|
||||
{
|
||||
pivot := lo;
|
||||
var i := lo+1;
|
||||
while i < hi
|
||||
invariant 0 <= lo <= pivot < i <= hi
|
||||
invariant forall k :: (0 <= k < lo || hi <= k < A.Length) ==> old(A[k]) == A[k]
|
||||
decreases hi - i
|
||||
{
|
||||
if A[i] < A[pivot] {
|
||||
var j := i-1;
|
||||
var tmp := A[i];
|
||||
A[i] := A[j];
|
||||
while pivot < j
|
||||
invariant forall k :: (0 <= k < lo || hi <= k < A.Length) ==> old(A[k]) == A[k]
|
||||
decreases j
|
||||
{
|
||||
A[j+1] := A[j];
|
||||
j := j-1;
|
||||
}
|
||||
A[pivot+1] := A[pivot];
|
||||
A[pivot] := tmp;
|
||||
pivot := pivot+1;
|
||||
}
|
||||
i := i+1;
|
||||
}
|
||||
}
|
||||
|
||||
method quicksort_between(A:array<int>, lo:int, hi:int)
|
||||
requires 0 <= lo <= hi <= A.Length
|
||||
ensures forall k :: (0 <= k < lo || hi <= k < A.Length) ==> old(A[k]) == A[k]
|
||||
modifies A
|
||||
decreases hi - lo
|
||||
{
|
||||
if lo+1 >= hi { return; }
|
||||
var pivot := partition(A, lo, hi);
|
||||
quicksort_between(A, lo, pivot);
|
||||
quicksort_between(A, pivot+1, hi);
|
||||
}
|
||||
|
||||
method quicksort(A:array<int>)
|
||||
modifies A
|
||||
{
|
||||
quicksort_between(A, 0, A.Length);
|
||||
}
|
||||
|
||||
method Main() {
|
||||
var A:array<int> := new int[7] [4,0,1,9,7,1,2];
|
||||
print "Before: ", A[0], A[1], A[2], A[3],
|
||||
A[4], A[5], A[6], "\n";
|
||||
quicksort(A);
|
||||
print "After: ", A[0], A[1], A[2], A[3],
|
||||
A[4], A[5], A[6], "\n";
|
||||
}
|
97
isabelle/2022/HSV_tasks_2022.thy
Normal file
97
isabelle/2022/HSV_tasks_2022.thy
Normal file
|
@ -0,0 +1,97 @@
|
|||
theory HSV_tasks_2022 imports Complex_Main begin
|
||||
|
||||
section {* Task 1: Full adders *}
|
||||
|
||||
fun fulladder :: "bool * bool * bool \<Rightarrow> bool * bool"
|
||||
where
|
||||
"fulladder (a,b,cin) = (
|
||||
let (tmp1, tmp2) = halfadder(a,b) in
|
||||
let (tmp3, s) = halfadder(cin,tmp2) in
|
||||
let cout = tmp1 | tmp3 in
|
||||
(cout, s))"
|
||||
|
||||
|
||||
section {* Task 2: Fifth powers *}
|
||||
|
||||
theorem "(n::nat) ^ 5 mod 10 = n mod 10"
|
||||
oops
|
||||
|
||||
section {* Task 3: Logic optimisation *}
|
||||
|
||||
(* Datatype for representing simple circuits. *)
|
||||
datatype "circuit" =
|
||||
NOT "circuit"
|
||||
| AND "circuit" "circuit"
|
||||
| OR "circuit" "circuit"
|
||||
| TRUE
|
||||
| FALSE
|
||||
| INPUT "int"
|
||||
|
||||
(* Simulates a circuit given a valuation for each input wire. *)
|
||||
fun simulate where
|
||||
"simulate (AND c1 c2) \<rho> = ((simulate c1 \<rho>) \<and> (simulate c2 \<rho>))"
|
||||
| "simulate (OR c1 c2) \<rho> = ((simulate c1 \<rho>) \<or> (simulate c2 \<rho>))"
|
||||
| "simulate (NOT c) \<rho> = (\<not> (simulate c \<rho>))"
|
||||
| "simulate TRUE \<rho> = True"
|
||||
| "simulate FALSE \<rho> = False"
|
||||
| "simulate (INPUT i) \<rho> = \<rho> i"
|
||||
|
||||
(* Equivalence between circuits. *)
|
||||
fun circuits_equiv (infix "\<sim>" 50) (* the "50" indicates the operator precedence *) where
|
||||
"c1 \<sim> c2 = (\<forall>\<rho>. simulate c1 \<rho> = simulate c2 \<rho>)"
|
||||
|
||||
(* An optimisation that exploits the following Boolean identities:
|
||||
`a | a = a`
|
||||
`a & a = a`
|
||||
*)
|
||||
fun opt_ident where
|
||||
"opt_ident (NOT c) = NOT (opt_ident c)"
|
||||
| "opt_ident (AND c1 c2) = (
|
||||
let c1' = opt_ident c1 in
|
||||
let c2' = opt_ident c2 in
|
||||
if c1' = c2' then c1' else AND c1' c2')"
|
||||
| "opt_ident (OR c1 c2) = (
|
||||
let c1' = opt_ident c1 in
|
||||
let c2' = opt_ident c2 in
|
||||
if c1' = c2' then c1' else OR c1' c2')"
|
||||
| "opt_ident TRUE = TRUE"
|
||||
| "opt_ident FALSE = FALSE"
|
||||
| "opt_ident (INPUT i) = INPUT i"
|
||||
|
||||
lemma (* test case *)
|
||||
"opt_ident (AND (INPUT 1) (OR (INPUT 1) (INPUT 1))) = INPUT 1"
|
||||
by eval
|
||||
|
||||
theorem opt_ident_is_sound: "opt_ident c \<sim> c"
|
||||
oops
|
||||
|
||||
fun area :: "circuit \<Rightarrow> nat" where
|
||||
"area (NOT c) = 1 + area c"
|
||||
| "area (AND c1 c2) = 1 + area c1 + area c2"
|
||||
| "area (OR c1 c2) = 1 + area c1 + area c2"
|
||||
| "area _ = 0"
|
||||
|
||||
section {* Task 4: More logic optimisation *}
|
||||
|
||||
lemma (* test case *)
|
||||
"opt_redundancy (AND (INPUT 1) (OR (INPUT 1) (INPUT 2)))
|
||||
= INPUT 1"
|
||||
(* by eval *) oops
|
||||
lemma (* test case *)
|
||||
"opt_redundancy (AND (AND (INPUT 1) (OR (INPUT 1) (INPUT 2)))
|
||||
(OR (AND (INPUT 1) (OR (INPUT 1) (INPUT 2))) (INPUT 2)))
|
||||
= INPUT 1"
|
||||
(* by eval *) oops
|
||||
lemma (* test case *)
|
||||
"opt_redundancy (AND (AND (INPUT 1) (OR (INPUT 1) (INPUT 2)))
|
||||
(OR (INPUT 2) (AND (INPUT 1) (OR (INPUT 1) (INPUT 2)))))
|
||||
= INPUT 1"
|
||||
(* by eval *) oops
|
||||
lemma (* test case *)
|
||||
"opt_redundancy (AND (AND (AND (INPUT 1) (OR (INPUT 1) (INPUT 2)))
|
||||
(OR (INPUT 2) (AND (INPUT 1) (OR (INPUT 1) (INPUT 2)))))
|
||||
(OR (INPUT 1) (INPUT 2)))
|
||||
= INPUT 1"
|
||||
(* by eval *) oops
|
||||
|
||||
end
|
BIN
isabelle/2022/isabelle_exercises_2022.pdf
Normal file
BIN
isabelle/2022/isabelle_exercises_2022.pdf
Normal file
Binary file not shown.
Loading…
Reference in a new issue