Skip to content

Commit 4897f58

Browse files
committed
saw-core-coq: fix and document the local translation state
5a37203 introduced a subtle bug in the translation of Cryptol modules to Coq. In this commit, it was decided that all 'TranslationState' should be restored when calling `withLocalLocalEnvironment'. Historically, 'withLocalLocalEnvironment' was only supposed to restore the value of the 'localEnvironment' field of the translation state, hence its name. The translation state has grown since, and includes many fields that ought to be restored after a local (sub-)term translation. Unfortunately, the translation state also contains fields that are meant to monotonically accumulate global data through the translation. Those fields are thus being erased incorrectly due to the changes made. This commit revert those changes, making it painfully explicit which fields of the state are to be restored or preserved so that future refactorings will have to make a decision. In the process, I also renamed the confusing 'localDeclarations' into 'topLevelDeclarations', since it captures declarations that appear in the processed file (so more local than the "global", "ambient" standard library declarations), but the name made it sound like these were more local.
1 parent 490f05a commit 4897f58

File tree

8 files changed

+177
-124
lines changed

8 files changed

+177
-124
lines changed

saw-core-coq/coq/generated/CryptolToCoq/CryptolPrimitivesForSAWCore.v

+5-2
Original file line numberDiff line numberDiff line change
@@ -583,8 +583,11 @@ Definition ecRotR : forall (m : @Num), forall (ix : Type), forall (a : Type), @P
583583
Definition ecCat : forall (m : @Num), forall (n : @Num), forall (a : Type), @seq m a -> @seq n a -> @seq (@tcAdd m n) a :=
584584
@finNumRec (fun (m : @Num) => forall (n : @Num), forall (a : Type), @seq m a -> @seq n a -> @seq (@tcAdd m n) a) (fun (m : @SAWCoreScaffolding.Nat) => @CryptolPrimitivesForSAWCore.Num_rect (fun (n : @Num) => forall (a : Type), @SAWCoreVectorsAsCoqVectors.Vec m a -> @seq n a -> @seq (@tcAdd (@TCNum m) n) a) (fun (n : @SAWCoreScaffolding.Nat) (a : Type) => @SAWCorePrelude.append m n a) (fun (a : Type) => @SAWCorePrelude.streamAppend a m)).
585585

586-
Definition ecSplitAt : forall (m : @Num), forall (n : @Num), forall (a : Type), @seq (@tcAdd m n) a -> prod (@seq m a) (@seq n a) :=
587-
@finNumRec (fun (m : @Num) => forall (n : @Num), forall (a : Type), @seq (@tcAdd m n) a -> prod (@seq m a) (@seq n a)) (fun (m : @SAWCoreScaffolding.Nat) => @CryptolPrimitivesForSAWCore.Num_rect (fun (n : @Num) => forall (a : Type), @seq (@tcAdd (@TCNum m) n) a -> prod (@SAWCoreVectorsAsCoqVectors.Vec m a) (@seq n a)) (fun (n : @SAWCoreScaffolding.Nat) (a : Type) (xs : @SAWCoreVectorsAsCoqVectors.Vec (@SAWCorePrelude.addNat m n) a) => pair (@SAWCorePrelude.take a m n xs) (@SAWCorePrelude.drop a m n xs)) (fun (a : Type) (xs : @SAWCorePrelude.Stream a) => pair (@SAWCorePrelude.streamTake a m xs) (@SAWCorePrelude.streamDrop a m xs))).
586+
Definition ecTake : forall (m : @Num), forall (n : @Num), forall (a : Type), @seq (@tcAdd m n) a -> @seq m a :=
587+
@CryptolPrimitivesForSAWCore.Num_rect (fun (m : @Num) => forall (n : @Num), forall (a : Type), @seq (@tcAdd m n) a -> @seq m a) (fun (m : @SAWCoreScaffolding.Nat) => @CryptolPrimitivesForSAWCore.Num_rect (fun (n : @Num) => forall (a : Type), @seq (@tcAdd (@TCNum m) n) a -> @SAWCoreVectorsAsCoqVectors.Vec m a) (fun (n : @SAWCoreScaffolding.Nat) (a : Type) (xs : @SAWCoreVectorsAsCoqVectors.Vec (@SAWCorePrelude.addNat m n) a) => @SAWCorePrelude.take a m n xs) (fun (a : Type) (xs : @SAWCorePrelude.Stream a) => @SAWCorePrelude.streamTake a m xs)) (@CryptolPrimitivesForSAWCore.Num_rect (fun (n : @Num) => forall (a : Type), @seq (@tcAdd (@TCInf) n) a -> @SAWCorePrelude.Stream a) (fun (n : @SAWCoreScaffolding.Nat) (a : Type) (xs : @SAWCorePrelude.Stream a) => xs) (fun (a : Type) (xs : @SAWCorePrelude.Stream a) => xs)).
588+
589+
Definition ecDrop : forall (m : @Num), forall (n : @Num), forall (a : Type), @seq (@tcAdd m n) a -> @seq n a :=
590+
@finNumRec (fun (m : @Num) => forall (n : @Num), forall (a : Type), @seq (@tcAdd m n) a -> @seq n a) (fun (m : @SAWCoreScaffolding.Nat) => @CryptolPrimitivesForSAWCore.Num_rect (fun (n : @Num) => forall (a : Type), @seq (@tcAdd (@TCNum m) n) a -> @seq n a) (fun (n : @SAWCoreScaffolding.Nat) (a : Type) (xs : @SAWCoreVectorsAsCoqVectors.Vec (@SAWCorePrelude.addNat m n) a) => @SAWCorePrelude.drop a m n xs) (fun (a : Type) (xs : @SAWCorePrelude.Stream a) => @SAWCorePrelude.streamDrop a m xs)).
588591

589592
Definition ecJoin : forall (m : @Num), forall (n : @Num), forall (a : Type), @seq m (@seq n a) -> @seq (@tcMul m n) a :=
590593
fun (m : @Num) => CryptolPrimitivesForSAWCore.Num_rect (fun (m1 : @Num) => forall (n : @Num), forall (a : Type), @seq m1 (@seq n a) -> @seq (@tcMul m1 n) a) (fun (m1 : @SAWCoreScaffolding.Nat) => @finNumRec (fun (n : @Num) => forall (a : Type), @SAWCoreVectorsAsCoqVectors.Vec m1 (@seq n a) -> @seq (@tcMul (@TCNum m1) n) a) (fun (n : @SAWCoreScaffolding.Nat) (a : Type) => @SAWCorePrelude.join m1 n a)) (@finNumRec (fun (n : @Num) => forall (a : Type), @SAWCorePrelude.Stream (@seq n a) -> @seq (@tcMul (@TCInf) n) a) (fun (n : @SAWCoreScaffolding.Nat) (a : Type) => @SAWCorePrelude.natCase (fun (n' : @SAWCoreScaffolding.Nat) => @SAWCorePrelude.Stream (@SAWCoreVectorsAsCoqVectors.Vec n' a) -> @seq (@SAWCorePrelude.if0Nat (@Num) n' (@TCNum 0) (@TCInf)) a) (fun (s : @SAWCorePrelude.Stream (@SAWCoreVectorsAsCoqVectors.Vec 0 a)) => @SAWCoreVectorsAsCoqVectors.EmptyVec a) (fun (n' : @SAWCoreScaffolding.Nat) (s : @SAWCorePrelude.Stream (@SAWCoreVectorsAsCoqVectors.Vec (@SAWCoreScaffolding.Succ n') a)) => @SAWCorePrelude.streamJoin a n' s) n)) m.

saw-core-coq/coq/generated/CryptolToCoq/SAWCorePrelude.v

+14-30
Large diffs are not rendered by default.

saw-core-coq/src/Verifier/SAW/Translation/Coq.hs

+10-2
Original file line numberDiff line numberDiff line change
@@ -113,7 +113,11 @@ Import ListNotations.
113113
translateTermAsDeclImports ::
114114
TranslationConfiguration -> Coq.Ident -> Term -> Either (TranslationError Term) (Doc ann)
115115
translateTermAsDeclImports configuration name t = do
116-
doc <- TermTranslation.translateDefDoc configuration Nothing [] name t
116+
doc <-
117+
TermTranslation.translateDefDoc
118+
configuration
119+
(TermTranslation.TranslationReader Nothing)
120+
[] name t
117121
return $ vcat [preamble configuration, hardline <> doc]
118122

119123
translateSAWModule :: TranslationConfiguration -> Module -> Doc ann
@@ -131,7 +135,11 @@ translateSAWModule configuration m =
131135
]
132136

133137
translateCryptolModule ::
134-
TranslationConfiguration -> [String] -> CryptolModule -> Either (TranslationError Term) (Doc ann)
138+
TranslationConfiguration ->
139+
-- | List of already translated global declarations
140+
[String] ->
141+
CryptolModule ->
142+
Either (TranslationError Term) (Doc ann)
135143
translateCryptolModule configuration globalDecls m =
136144
let decls = CryptolModuleTranslation.translateCryptolModule
137145
configuration

saw-core-coq/src/Verifier/SAW/Translation/Coq/CryptolModule.hs

+20-13
Original file line numberDiff line numberDiff line change
@@ -8,10 +8,10 @@ import Control.Monad (forM)
88
import Control.Monad.State (modify)
99
import qualified Data.Map as Map
1010

11-
import Cryptol.ModuleSystem.Name
12-
import Cryptol.Utils.Ident
11+
import Cryptol.ModuleSystem.Name (Name, nameIdent)
12+
import Cryptol.Utils.Ident (unpackIdent)
1313
import qualified Language.Coq.AST as Coq
14-
import Verifier.SAW.Term.Functor
14+
import Verifier.SAW.Term.Functor (Term)
1515
import Verifier.SAW.Translation.Coq.Monad
1616
import qualified Verifier.SAW.Translation.Coq.Term as TermTranslation
1717
import Verifier.SAW.TypedTerm
@@ -23,21 +23,28 @@ translateTypedTermMap tm = forM (Map.assocs tm) translateAndRegisterEntry
2323
translateAndRegisterEntry (name, symbol) = do
2424
let t = ttTerm symbol
2525
let nameStr = unpackIdent (nameIdent name)
26-
term <- TermTranslation.withLocalLocalEnvironment $ do
26+
term <- TermTranslation.withLocalTranslationState $ do
2727
modify $ set TermTranslation.localEnvironment [nameStr]
2828
TermTranslation.translateTerm t
2929
let decl = TermTranslation.mkDefinition nameStr term
3030
modify $ over TermTranslation.globalDeclarations (nameStr :)
3131
return decl
3232

33+
-- | Translates a Cryptol Module into a list of Coq declarations. This is done
34+
-- by going through the term map corresponding to the module, translating all
35+
-- terms, and accumulating the translated declarations of all top-level
36+
-- declarations encountered.
3337
translateCryptolModule ::
34-
TranslationConfiguration -> [String] -> CryptolModule -> Either (TranslationError Term) [Coq.Decl]
38+
TranslationConfiguration ->
39+
-- | List of already translated global declarations
40+
[String] ->
41+
CryptolModule ->
42+
Either (TranslationError Term) [Coq.Decl]
3543
translateCryptolModule configuration globalDecls (CryptolModule _ tm) =
36-
case TermTranslation.runTermTranslationMonad
37-
configuration
38-
Nothing
39-
globalDecls
40-
[]
41-
(translateTypedTermMap tm) of
42-
Left e -> Left e
43-
Right (_, st) -> Right (reverse (view TermTranslation.localDeclarations st))
44+
reverse . view TermTranslation.topLevelDeclarations . snd
45+
<$> TermTranslation.runTermTranslationMonad
46+
configuration
47+
(TermTranslation.TranslationReader Nothing) -- TODO: this should be Just no?
48+
globalDecls
49+
[]
50+
(translateTypedTermMap tm)

saw-core-coq/src/Verifier/SAW/Translation/Coq/Monad.hs

+20-6
Original file line numberDiff line numberDiff line change
@@ -17,6 +17,7 @@ module Verifier.SAW.Translation.Coq.Monad
1717
, TranslationConfigurationMonad
1818
, TranslationMonad
1919
, TranslationError(..)
20+
, WithTranslationConfiguration(..)
2021
, runTranslationMonad
2122
) where
2223

@@ -77,19 +78,32 @@ data TranslationConfiguration = TranslationConfiguration
7778
-- - SAWCoreVectorsAsSSReflectSeqs
7879
}
7980

80-
type TranslationConfigurationMonad m =
81-
( MonadReader TranslationConfiguration m
81+
-- | The functional dependency of 'MonadReader' makes it not compositional, so
82+
-- we have to jam together different structures that want to be in the 'Reader'
83+
-- into a single datatype. This type allows adding extra configuration on top
84+
-- of the translation configuration.
85+
data WithTranslationConfiguration r = WithTranslationConfiguration
86+
{ translationConfiguration :: TranslationConfiguration
87+
, otherConfiguration :: r
88+
}
89+
90+
-- | Some computations will rely solely on access to the configuration, so we
91+
-- provide it separately.
92+
type TranslationConfigurationMonad r m =
93+
( MonadReader (WithTranslationConfiguration r) m
8294
)
8395

84-
type TranslationMonad s m =
96+
type TranslationMonad r s m =
8597
( Except.MonadError (TranslationError Term) m
86-
, TranslationConfigurationMonad m
98+
, TranslationConfigurationMonad r m
8799
, MonadState s m
88100
)
89101

90102
runTranslationMonad ::
91103
TranslationConfiguration ->
104+
r ->
92105
s ->
93-
(forall m. TranslationMonad s m => m a) ->
106+
(forall m. TranslationMonad r s m => m a) ->
94107
Either (TranslationError Term) (a, s)
95-
runTranslationMonad configuration state m = runStateT (runReaderT m configuration) state
108+
runTranslationMonad configuration r s m =
109+
runStateT (runReaderT m (WithTranslationConfiguration configuration r)) s

saw-core-coq/src/Verifier/SAW/Translation/Coq/SAWModule.hs

+16-18
Original file line numberDiff line numberDiff line change
@@ -24,10 +24,8 @@ Portability : portable
2424

2525
module Verifier.SAW.Translation.Coq.SAWModule where
2626

27-
import Control.Lens (makeLenses, view)
2827
import qualified Control.Monad.Except as Except
2928
import Control.Monad.Reader hiding (fail)
30-
import Control.Monad.State hiding (fail)
3129
import Prelude hiding (fail)
3230
import Prettyprinter (Doc)
3331

@@ -36,26 +34,22 @@ import qualified Language.Coq.Pretty as Coq
3634
import Verifier.SAW.Module
3735
import Verifier.SAW.SharedTerm
3836
import Verifier.SAW.Term.Functor
39-
import Verifier.SAW.Translation.Coq.Monad
37+
import qualified Verifier.SAW.Translation.Coq.Monad as M
4038
import Verifier.SAW.Translation.Coq.SpecialTreatment
4139
import qualified Verifier.SAW.Translation.Coq.Term as TermTranslation
40+
import Verifier.SAW.Translation.Coq.Monad
4241

4342
-- import Debug.Trace
4443

45-
data ModuleTranslationState = ModuleTranslationState
46-
{ _currentModule :: Maybe ModuleName
47-
}
48-
deriving (Show)
49-
makeLenses ''ModuleTranslationState
50-
51-
type ModuleTranslationMonad m = TranslationMonad ModuleTranslationState m
44+
type ModuleTranslationMonad m =
45+
M.TranslationMonad TermTranslation.TranslationReader () m
5246

5347
runModuleTranslationMonad ::
54-
TranslationConfiguration -> Maybe ModuleName ->
48+
M.TranslationConfiguration -> Maybe ModuleName ->
5549
(forall m. ModuleTranslationMonad m => m a) ->
56-
Either (TranslationError Term) (a, ModuleTranslationState)
57-
runModuleTranslationMonad configuration modname =
58-
runTranslationMonad configuration (ModuleTranslationState modname)
50+
Either (M.TranslationError Term) (a, ())
51+
runModuleTranslationMonad configuration modName =
52+
M.runTranslationMonad configuration (TermTranslation.TranslationReader modName) ()
5953

6054
dropPi :: Coq.Term -> Coq.Term
6155
dropPi (Coq.Pi (_ : t) r) = Coq.Pi t r
@@ -162,15 +156,19 @@ liftTermTranslationMonad ::
162156
(forall n. TermTranslation.TermTranslationMonad n => n a) ->
163157
(forall m. ModuleTranslationMonad m => m a)
164158
liftTermTranslationMonad n = do
165-
configuration <- ask
166-
cur_mod <- view currentModule <$> get
167-
let r = TermTranslation.runTermTranslationMonad configuration cur_mod [] [] n
159+
configuration <- asks translationConfiguration
160+
configReader <- asks otherConfiguration
161+
let r = TermTranslation.runTermTranslationMonad configuration configReader [] [] n
168162
case r of
169163
Left e -> Except.throwError e
170164
Right (a, _) -> do
171165
return a
172166

173-
translateDecl :: TranslationConfiguration -> Maybe ModuleName -> ModuleDecl -> Doc ann
167+
translateDecl ::
168+
M.TranslationConfiguration ->
169+
Maybe ModuleName ->
170+
ModuleDecl ->
171+
Doc ann
174172
translateDecl configuration modname decl =
175173
case decl of
176174
TypeDecl td -> do

saw-core-coq/src/Verifier/SAW/Translation/Coq/SpecialTreatment.hs

+3-3
Original file line numberDiff line numberDiff line change
@@ -25,7 +25,7 @@ Portability : portable
2525
module Verifier.SAW.Translation.Coq.SpecialTreatment where
2626

2727
import Control.Lens (_1, _2, over)
28-
import Control.Monad.Reader (ask)
28+
import Control.Monad.Reader (asks)
2929
import qualified Data.Map as Map
3030
import Data.String.Interpolate (i)
3131
import qualified Data.Text as Text
@@ -70,10 +70,10 @@ translateModuleName mn =
7070
Map.findWithDefault mn mn moduleRenamingMap
7171

7272
findSpecialTreatment ::
73-
TranslationConfigurationMonad m =>
73+
TranslationConfigurationMonad r m =>
7474
Ident -> m IdentSpecialTreatment
7575
findSpecialTreatment ident = do
76-
configuration <- ask
76+
configuration <- asks translationConfiguration
7777
let moduleMap =
7878
Map.findWithDefault Map.empty (identModule ident) (specialTreatmentMap configuration)
7979
let defaultTreatment =

0 commit comments

Comments
 (0)