Theory Prelim
theory Prelim
imports "Fresh_Identifiers.Fresh_String" "Bounded_Deducibility_Security.Trivia"
begin
subsection ‹The basic types›
type_synonym string = String.literal
definition "emptyStr = STR ''''"
type_synonym phase = nat
abbreviation "noPH ≡ (0::nat)" abbreviation "setPH ≡ Suc noPH" abbreviation "subPH ≡ Suc setPH"
abbreviation "bidPH ≡ Suc subPH" abbreviation "revPH ≡ Suc bidPH" abbreviation "disPH ≡ Suc revPH"
abbreviation "notifPH ≡ Suc disPH" abbreviation "closedPH ≡ Suc notifPH"
fun SucPH where
"SucPH ph = (if ph = closedPH then closedPH else Suc ph)"
datatype user = User string string string
fun nameUser where "nameUser (User name info email) = name"
fun infoUser where "infoUser (User name info email) = info"
fun emailUser where "emailUser (User name info email) = email"
definition "emptyUser ≡ User emptyStr emptyStr emptyStr"
typedecl raw_data
code_printing type_constructor raw_data ⇀ (Scala) "java.io.File"
datatype pcontent = NoPContent | PContent raw_data