src/HOL/Induct/QuoDataType.thy
changeset 58305 57752a91eec4
parent 49834 b27bbb021df1
child 58310 91ea607a34d8
equal deleted inserted replaced
58304:acc2f1801acc 58305:57752a91eec4
     8 theory QuoDataType imports Main begin
     8 theory QuoDataType imports Main begin
     9 
     9 
    10 subsection{*Defining the Free Algebra*}
    10 subsection{*Defining the Free Algebra*}
    11 
    11 
    12 text{*Messages with encryption and decryption as free constructors.*}
    12 text{*Messages with encryption and decryption as free constructors.*}
    13 datatype
    13 datatype_new
    14      freemsg = NONCE  nat
    14      freemsg = NONCE  nat
    15              | MPAIR  freemsg freemsg
    15              | MPAIR  freemsg freemsg
    16              | CRYPT  nat freemsg  
    16              | CRYPT  nat freemsg  
    17              | DECRYPT  nat freemsg
    17              | DECRYPT  nat freemsg
    18 
    18