mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 06:52:19 +03:00
[ papers ] A bit of tidying up and clarification
This commit is contained in:
parent
d25db9e353
commit
cfe18b3d01
@ -391,27 +391,30 @@ parameters (Lbls, Sts : Type)
|
||||
agSearch : {f : _} -> MC f -> MC (AlwaysGlobal f)
|
||||
agSearch g = auSearch g (g `mcAND'` isCompleted)
|
||||
|
||||
|
||||
------------------------------------------------------------------------
|
||||
-- Proof search example
|
||||
|
||||
||| This CT is a model of composing the `HiHorse` and `LoRoad` programs.
|
||||
public export
|
||||
Tree : CT ((), ()) Nat
|
||||
Tree = model ((), ()) Nat (HiHorse `pComp` LoRoad) 0
|
||||
tree : CT ((), ()) Nat
|
||||
tree = model ((), ()) Nat (HiHorse `pComp` LoRoad) 0
|
||||
|
||||
||| Prove that there exists a path where `HiHorse || LoRoad`'s state reaches 10.
|
||||
||| A half-decider for proving that there exists a path where the shared
|
||||
||| `HiHorse || LoRoad` state reaches 10.
|
||||
public export
|
||||
reaches10 : ? -- HDec (ExistsFinally [...])
|
||||
reaches10 =
|
||||
efSearch ((), ()) Nat
|
||||
(now ((), ()) Nat
|
||||
(\ st, _ => fromDec $ decEq st 10)) Tree 20
|
||||
(\st, _ => fromDec $ decEq st 10)) tree 20
|
||||
|
||||
||| Prove that the shared state of `HiHorse || LoRoad` reaches 10, using the
|
||||
||| previously defined half-decider.
|
||||
export
|
||||
r10Proof : Models ((), ()) Nat
|
||||
Tree
|
||||
CTL.tree
|
||||
(ExistsFinally ((), ()) Nat
|
||||
(Guarded ((), ()) Nat (\ st, _ => st === 10)))
|
||||
r10Proof = diModels ((), ()) Nat (reaches10.evidence Oh)
|
||||
|
||||
|
||||
|
@ -421,10 +421,10 @@ dekkers : Diagram (GCL State, GCL State) State
|
||||
dekkers = (gclToDiag _ dekkers1) `pComp` (gclToDiag _ dekkers2)
|
||||
|
||||
||| An attempt at finding a violation of Mutual Exclusion.
|
||||
||| THIS WILL NOT WORK due to the lack of fairness in the unfolding of the
|
||||
||| traces. Dekker's algorithm requires fair scheduling in order to be correct,
|
||||
||| but since we don't have that, we cannot find a proof that no violations of
|
||||
||| mutex exist.
|
||||
||| THIS WILL NOT FIND A PROOF due to the lack of fairness in the unfolding of
|
||||
||| the traces. Dekker's algorithm requires fair scheduling in order to be
|
||||
||| correct, but since we don't have that, we cannot find a proof that no
|
||||
||| violations of mutex exist.
|
||||
|||
|
||||
||| /!\ Trying to evaluate this did not finish after 10 minutes /!\
|
||||
public export
|
||||
|
Loading…
Reference in New Issue
Block a user