Theory SAT_Plan_Base
theory SAT_Plan_Base
imports "List-Index.List_Index"
"Propositional_Proof_Systems.Formulas"
"STRIPS_Semantics"
"Map_Supplement" "List_Supplement"
"CNF_Semantics_Supplement" "CNF_Supplement"
begin
hide_const (open) Orderings.bot_class.bot
no_notation Orderings.bot_class.bot (‹⊥›)
hide_const (open) Transitive_Closure.trancl
unbundle no trancl_syntax
hide_const (open) Relation.converse
no_notation Relation.converse (‹(‹notation=‹postfix -1››_¯)› [1000] 999)
section "The Basic SATPlan Encoding"
text ‹ We now move on to the formalization of the basic SATPlan encoding (see
\autoref{def:basic-sat-plan-encoding-strips-problem}).
The two major results that we will obtain here are the soundness and completeness result outlined
in \autoref{thm:soundness-and-completeness-satplan-base} in
\autoref{sub:soundness-completeness-satplan}.
Let in the following ‹Φ ≡ encode_to_sat Π t› denote the SATPlan encoding for a STRIPS problem ‹Π›
and makespan ‹t›. Let \<^term>‹k < t› and ‹I ≡ (Π)⇩I› be the initial state of ‹Π›, ‹G ≡ (Π)⇩G› be
its goal state, ‹𝒱 ≡ (Π)⇩𝒱› its variable set, and ‹𝒪 ≡ (Π)⇩𝒪› its operator set. ›
subsection "Encoding Function Definitions"
text ‹ Since the SATPlan encoding uses propositional variables for both operators and state
variables of the problem as well as time points, we define a datatype using separate constructors
---\<^term>‹State k n› for state variables resp. \<^term>‹Operator k n› for operator activation---to
facilitate case distinction.
The natural number values store the time index resp. the indexes of the variable or operator
within their lists in the problem representation.
% TODO Note on why formulas are used instead of CNF (simple representation and good basis; e.g.
% export to cnf lists using CNF_Formulas.cnf_lists) ›