doc-src/IsarRef/Thy/HOL_Specific.thy
changeset 42912 a5bbc11474f9
parent 42911 6891e8a8d748
child 42913 68bc69bdce88
equal deleted inserted replaced
42911:6891e8a8d748 42912:a5bbc11474f9
   169   "}
   169   "}
   170 
   170 
   171   \begin{description}
   171   \begin{description}
   172 
   172 
   173   \item @{command (HOL) "primrec"} defines primitive recursive
   173   \item @{command (HOL) "primrec"} defines primitive recursive
   174   functions over datatypes, see also \cite{isabelle-HOL}.
   174   functions over datatypes (see also @{command_ref (HOL) datatype} and
       
   175   @{command_ref (HOL) rep_datatype}).  The given @{text equations}
       
   176   specify reduction rules that are produced by instantiating the
       
   177   generic combinator for primitive recursion that is available for
       
   178   each datatype.
       
   179 
       
   180   Each equation needs to be of the form:
       
   181 
       
   182   @{text [display] "f x\<^sub>1 \<dots> x\<^sub>m (C y\<^sub>1 \<dots> y\<^sub>k) z\<^sub>1 \<dots> z\<^sub>n = rhs"}
       
   183 
       
   184   such that @{text C} is a datatype constructor, @{text rhs} contains
       
   185   only the free variables on the left-hand side (or from the context),
       
   186   and all recursive occurrences of @{text "f"} in @{text "rhs"} are of
       
   187   the form @{text "f \<dots> y\<^sub>i \<dots>"} for some @{text i}.  At most one
       
   188   reduction rule for each constructor can be given.  The order does
       
   189   not matter.  For missing constructors, the function is defined to
       
   190   return a default value, but this equation is made difficult to
       
   191   access for users.
       
   192 
       
   193   The reduction rules are declared as @{attribute simp} by default,
       
   194   which enables standard proof methods like @{method simp} and
       
   195   @{method auto} to normalize expressions of @{text "f"} applied to
       
   196   datatype constructions, by simulating symbolic computation via
       
   197   rewriting.
   175 
   198 
   176   \item @{command (HOL) "function"} defines functions by general
   199   \item @{command (HOL) "function"} defines functions by general
   177   wellfounded recursion. A detailed description with examples can be
   200   wellfounded recursion. A detailed description with examples can be
   178   found in \cite{isabelle-function}. The function is specified by a
   201   found in \cite{isabelle-function}. The function is specified by a
   179   set of (possibly conditional) recursive equations with arbitrary
   202   set of (possibly conditional) recursive equations with arbitrary
   198   the induction principle is established.
   221   the induction principle is established.
   199 
   222 
   200   \end{description}
   223   \end{description}
   201 
   224 
   202   Recursive definitions introduced by the @{command (HOL) "function"}
   225   Recursive definitions introduced by the @{command (HOL) "function"}
   203   command accommodate
   226   command accommodate reasoning by induction (cf.\ @{method induct}):
   204   reasoning by induction (cf.\ \secref{sec:cases-induct}): rule @{text
   227   rule @{text "f.induct"} refers to a specific induction rule, with
   205   "c.induct"} (where @{text c} is the name of the function definition)
   228   parameters named according to the user-specified equations. Cases
   206   refers to a specific induction rule, with parameters named according
   229   are numbered starting from 1.  For @{command (HOL) "primrec"}, the
   207   to the user-specified equations. Cases are numbered (starting from 1).
   230   induction principle coincides with structural recursion on the
   208 
   231   datatype where the recursion is carried out.
   209   For @{command (HOL) "primrec"}, the induction principle coincides
       
   210   with structural recursion on the datatype the recursion is carried
       
   211   out.
       
   212 
   232 
   213   The equations provided by these packages may be referred later as
   233   The equations provided by these packages may be referred later as
   214   theorem list @{text "f.simps"}, where @{text f} is the (collective)
   234   theorem list @{text "f.simps"}, where @{text f} is the (collective)
   215   name of the functions defined.  Individual equations may be named
   235   name of the functions defined.  Individual equations may be named
   216   explicitly as well.
   236   explicitly as well.
   234   introduction rules for the domain predicate. While mostly not
   254   introduction rules for the domain predicate. While mostly not
   235   needed, they can be helpful in some proofs about partial functions.
   255   needed, they can be helpful in some proofs about partial functions.
   236 
   256 
   237   \end{description}
   257   \end{description}
   238 *}
   258 *}
       
   259 
       
   260 subsubsection {* Example: evaluation of expressions *}
       
   261 
       
   262 text {* Subsequently, we define mutual datatypes for arithmetic and
       
   263   boolean expressions, and use @{command primrec} for evaluation
       
   264   functions that follow the same recursive structure. *}
       
   265 
       
   266 datatype 'a aexp =
       
   267     IF "'a bexp"  "'a aexp"  "'a aexp"
       
   268   | Sum "'a aexp"  "'a aexp"
       
   269   | Diff "'a aexp"  "'a aexp"
       
   270   | Var 'a
       
   271   | Num nat
       
   272 and 'a bexp =
       
   273     Less "'a aexp"  "'a aexp"
       
   274   | And "'a bexp"  "'a bexp"
       
   275   | Neg "'a bexp"
       
   276 
       
   277 
       
   278 text {* \medskip Evaluation of arithmetic and boolean expressions *}
       
   279 
       
   280 primrec evala :: "('a \<Rightarrow> nat) \<Rightarrow> 'a aexp \<Rightarrow> nat"
       
   281   and evalb :: "('a \<Rightarrow> nat) \<Rightarrow> 'a bexp \<Rightarrow> bool"
       
   282 where
       
   283   "evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)"
       
   284 | "evala env (Sum a1 a2) = evala env a1 + evala env a2"
       
   285 | "evala env (Diff a1 a2) = evala env a1 - evala env a2"
       
   286 | "evala env (Var v) = env v"
       
   287 | "evala env (Num n) = n"
       
   288 | "evalb env (Less a1 a2) = (evala env a1 < evala env a2)"
       
   289 | "evalb env (And b1 b2) = (evalb env b1 \<and> evalb env b2)"
       
   290 | "evalb env (Neg b) = (\<not> evalb env b)"
       
   291 
       
   292 text {* Since the value of an expression depends on the value of its
       
   293   variables, the functions @{const evala} and @{const evalb} take an
       
   294   additional parameter, an \emph{environment} that maps variables to
       
   295   their values.
       
   296 
       
   297   \medskip Substitution on expressions can be defined similarly.  The
       
   298   mapping @{text f} of type @{typ "'a \<Rightarrow> 'a aexp"} given as a
       
   299   parameter is lifted canonically on the types @{typ "'a aexp"} and
       
   300   @{typ "'a bexp"}, respectively.
       
   301 *}
       
   302 
       
   303 primrec substa :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a aexp \<Rightarrow> 'b aexp"
       
   304   and substb :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a bexp \<Rightarrow> 'b bexp"
       
   305 where
       
   306   "substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)"
       
   307 | "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)"
       
   308 | "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)"
       
   309 | "substa f (Var v) = f v"
       
   310 | "substa f (Num n) = Num n"
       
   311 | "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)"
       
   312 | "substb f (And b1 b2) = And (substb f b1) (substb f b2)"
       
   313 | "substb f (Neg b) = Neg (substb f b)"
       
   314 
       
   315 text {* In textbooks about semantics one often finds substitution
       
   316   theorems, which express the relationship between substitution and
       
   317   evaluation.  For @{typ "'a aexp"} and @{typ "'a bexp"}, we can prove
       
   318   such a theorem by mutual induction, followed by simplification.
       
   319 *}
       
   320 
       
   321 lemma subst_one:
       
   322   "evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a"
       
   323   "evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b"
       
   324   by (induct a and b) simp_all
       
   325 
       
   326 lemma subst_all:
       
   327   "evala env (substa s a) = evala (\<lambda>x. evala env (s x)) a"
       
   328   "evalb env (substb s b) = evalb (\<lambda>x. evala env (s x)) b"
       
   329   by (induct a and b) simp_all
       
   330 
       
   331 
       
   332 subsubsection {* Example: a substitution function for terms *}
       
   333 
       
   334 text {* Functions on datatypes with nested recursion are also defined
       
   335   by mutual primitive recursion. *}
       
   336 
       
   337 datatype ('a, 'b) "term" = Var 'a | App 'b "('a, 'b) term list"
       
   338 
       
   339 text {* A substitution function on type @{typ "('a, 'b) term"} can be
       
   340   defined as follows, by working simultaneously on @{typ "('a, 'b)
       
   341   term list"}: *}
       
   342 
       
   343 primrec subst_term :: "('a \<Rightarrow> ('a, 'b) term) \<Rightarrow> ('a, 'b) term \<Rightarrow> ('a, 'b) term" and
       
   344   subst_term_list :: "('a \<Rightarrow> ('a, 'b) term) \<Rightarrow> ('a, 'b) term list \<Rightarrow> ('a, 'b) term list"
       
   345 where
       
   346   "subst_term f (Var a) = f a"
       
   347 | "subst_term f (App b ts) = App b (subst_term_list f ts)"
       
   348 | "subst_term_list f [] = []"
       
   349 | "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts"
       
   350 
       
   351 text {* The recursion scheme follows the structure of the unfolded
       
   352   definition of type @{typ "('a, 'b) term"}.  To prove properties of this
       
   353   substitution function, mutual induction is needed:
       
   354 *}
       
   355 
       
   356 lemma "subst_term (subst_term f1 \<circ> f2) t = subst_term f1 (subst_term f2 t)" and
       
   357   "subst_term_list (subst_term f1 \<circ> f2) ts = subst_term_list f1 (subst_term_list f2 ts)"
       
   358   by (induct t and ts) simp_all
       
   359 
       
   360 
       
   361 subsubsection {* Example: a map function for infinitely branching trees *}
       
   362 
       
   363 text {* Defining functions on infinitely branching datatypes by
       
   364   primitive recursion is just as easy.
       
   365 *}
       
   366 
       
   367 datatype 'a tree = Atom 'a | Branch "nat \<Rightarrow> 'a tree"
       
   368 
       
   369 primrec map_tree :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a tree \<Rightarrow> 'b tree"
       
   370 where
       
   371   "map_tree f (Atom a) = Atom (f a)"
       
   372 | "map_tree f (Branch ts) = Branch (\<lambda>x. map_tree f (ts x))"
       
   373 
       
   374 text {* Note that all occurrences of functions such as @{text ts}
       
   375   above must be applied to an argument.  In particular, @{term
       
   376   "map_tree f \<circ> ts"} is not allowed here. *}
       
   377 
       
   378 text {* Here is a simple composition lemma for @{term map_tree}: *}
       
   379 
       
   380 lemma "map_tree g (map_tree f t) = map_tree (g \<circ> f) t"
       
   381   by (induct t) simp_all
   239 
   382 
   240 
   383 
   241 subsection {* Proof methods related to recursive definitions *}
   384 subsection {* Proof methods related to recursive definitions *}
   242 
   385 
   243 text {*
   386 text {*