Theory Generic_ADS_Construction
theory Generic_ADS_Construction imports
"Merkle_Interface"
"HOL-Library.BNF_Axiomatization"
begin
section ‹Generic construction of authenticated data structures›
subsection ‹Functors›
subsubsection ‹Source functor›
text ‹
We want to allow ADSs of arbitrary ADTs, which we call "source trees". The ADTs we are interested in can
in general be represented as the least fixpoints of some bounded natural (bi-)functor (BNF) ‹('a, 'b) F›, where
@{typ 'a} is the type of "source" data, and @{typ 'b} is a recursion "handle".
However, Isabelle's type system does not support higher kinds, necessary to parameterize our definitions
over functors.
Instead, we first develop a general theory of ADSs over an arbitrary, but fixed functor,
and its least fixpoint. We show that the theory is compositional, in that the functor's least fixed point
can then be reused as the "source" data of another functor.
We start by defining the arbitrary fixed functor, its fixpoints, and showing how composition can be
done. A higher-level explanation is found in the paper.
›
bnf_axiomatization ('a, 'b) F [wits: "'a ⇒ ('a, 'b) F"]
context notes [[typedef_overloaded]] begin