Abstract
We introduce Jinja, a Java-like programming language with a formal semantics designed to exhibit core features of the Java language architecture. Jinja is a compromise between realism of the language and tractability and clarity of the formal semantics. The following aspects are formalised: a big and a small step operational semantics for Jinja and a proof of their equivalence; a type system and a definite initialisation analysis; a type safety proof of the small step semantics; a virtual machine (JVM), its operational semantics and its type system; a type safety proof for the JVM; a bytecode verifier, i.e. data flow analyser for the JVM; a correctness proof of the bytecode verifier w.r.t. the type system; a compiler and a proof that it preserves semantics and well-typedness. The emphasis of this work is not on particular language features but on providing a unified model of the source language, the virtual machine and the compiler. The whole development has been carried out in the theorem prover Isabelle/HOL.
License
Topics
Related publications
- Klein, G., & Nipkow, T. (2006). A machine-checked model for a Java-like language, virtual machine, and compiler. ACM Transactions on Programming Languages and Systems, 28(4), 619–695. https://doi.org/10.1145/1146809.1146811
- Alternate versions
Session Jinja
- Auxiliary
- Type
- Decl
- TypeRel
- Value
- Objects
- Exceptions
- Expr
- State
- BigStep
- SmallStep
- SystemClasses
- WellForm
- WWellForm
- Equivalence
- WellType
- WellTypeRT
- DefAss
- Conform
- Progress
- JWellForm
- TypeSafe
- Annotate
- Examples
- execute_Bigstep
- execute_WellType
- JVMState
- JVMInstructions
- JVMExecInstr
- JVMExceptions
- JVMExec
- JVMDefensive
- JVMListExample
- Semilat
- Err
- Opt
- Product
- Listn
- Semilattices
- Typing_Framework_1
- SemilatAlg
- Typing_Framework_err
- Kildall_1
- Kildall_2
- LBVSpec
- LBVCorrect
- LBVComplete
- Abstract_BV
- SemiType
- JVM_SemiType
- Effect
- EffectMono
- BVSpec
- TF_JVM
- Typing_Framework_2
- BVExec
- LBVJVM
- BVConform
- BVSpecTypeSafe
- BVNoTypeError
- BVExample
- J1
- J1WellForm
- PCompiler
- Hidden
- Compiler1
- Correctness1
- Compiler2
- Correctness2
- Compiler
- TypeComp
- Jinja