Theory Fixrec_ex
section ‹Fixrec package examples›
theory Fixrec_ex
imports HOLCF
begin
subsection ‹Basic ‹fixrec› examples›
text ‹
Fixrec patterns can mention any constructor defined by the domain
package, as well as any of the following built-in constructors:
Pair, spair, sinl, sinr, up, ONE, TT, FF.
›
text ‹Typical usage is with lazy constructors.›
fixrec down :: "'a u → 'a"
where "down⋅(up⋅x) = x"
text ‹With strict constructors, rewrite rules may require side conditions.›
fixrec from_sinl :: "'a ⊕ 'b → 'a"
where "x ≠ ⊥ ⟹ from_sinl⋅(sinl⋅x) = x"
text ‹Lifting can turn a strict constructor into a lazy one.›
fixrec from_sinl_up :: "'a u ⊕ 'b → 'a"
where "from_sinl_up⋅(sinl⋅(up⋅x)) = x"
text ‹Fixrec also works with the HOL pair constructor.›
fixrec down2 :: "'a u × 'b u → 'a × 'b"
where "down2⋅(up⋅x, up⋅y) = (x, y)"
subsection ‹Examples using ‹fixrec_simp››
text ‹A type of lazy lists.›
domain 'a llist = lNil | lCons (lazy 'a) (lazy "'a llist")
text ‹A zip function for lazy lists.›
text ‹Notice that the patterns are not exhaustive.›
fixrec
lzip :: "'a llist → 'b llist → ('a × 'b) llist"
where
"lzip⋅(lCons⋅x⋅xs)⋅(lCons⋅y⋅ys) = lCons⋅(x, y)⋅(lzip⋅xs⋅ys)"
| "lzip⋅lNil⋅lNil = lNil"
text ‹‹fixrec_simp› is useful for producing strictness theorems.›
text ‹Note that pattern matching is done in left-to-right order.›
lemma lzip_stricts [simp]:
"lzip⋅⊥⋅ys = ⊥"
"lzip⋅lNil⋅⊥ = ⊥"
"lzip⋅(lCons⋅x⋅xs)⋅⊥ = ⊥"
by fixrec_simp+
text ‹‹fixrec_simp› can also produce rules for missing cases.›
lemma lzip_undefs [simp]:
"lzip⋅lNil⋅(lCons⋅y⋅ys) = ⊥"
"lzip⋅(lCons⋅x⋅xs)⋅lNil = ⊥"
by fixrec_simp+
subsection ‹Pattern matching with bottoms›
text ‹
As an alternative to using ‹fixrec_simp›, it is also possible
to use bottom as a constructor pattern. When using a bottom
pattern, the right-hand-side must also be bottom; otherwise, ‹fixrec› will not be able to prove the equation.
›
fixrec
from_sinr_up :: "'a ⊕ 'b⇩⊥ → 'b"
where
"from_sinr_up⋅⊥ = ⊥"
| "from_sinr_up⋅(sinr⋅(up⋅x)) = x"
text ‹
If the function is already strict in that argument, then the bottom
pattern does not change the meaning of the function. For example,
in the definition of \<^term>‹from_sinr_up›, the first equation is
actually redundant, and could have been proven separately by
‹fixrec_simp›.
›
text ‹
A bottom pattern can also be used to make a function strict in a
certain argument, similar to a bang-pattern in Haskell.
›
fixrec
seq :: "'a → 'b → 'b"
where
"seq⋅⊥⋅y = ⊥"
| "x ≠ ⊥ ⟹ seq⋅x⋅y = y"
subsection ‹Skipping proofs of rewrite rules›
text ‹Another zip function for lazy lists.›
text ‹
Notice that this version has overlapping patterns.
The second equation cannot be proved as a theorem
because it only applies when the first pattern fails.
›
fixrec
lzip2 :: "'a llist → 'b llist → ('a × 'b) llist"
where
"lzip2⋅(lCons⋅x⋅xs)⋅(lCons⋅y⋅ys) = lCons⋅(x, y)⋅(lzip2⋅xs⋅ys)"
| (unchecked) "lzip2⋅xs⋅ys = lNil"
text ‹
Usually fixrec tries to prove all equations as theorems.
The "unchecked" option overrides this behavior, so fixrec
does not attempt to prove that particular equation.
›
text ‹Simp rules can be generated later using ‹fixrec_simp›.›
lemma lzip2_simps [simp]:
"lzip2⋅(lCons⋅x⋅xs)⋅lNil = lNil"
"lzip2⋅lNil⋅(lCons⋅y⋅ys) = lNil"
"lzip2⋅lNil⋅lNil = lNil"
by fixrec_simp+
lemma lzip2_stricts [simp]:
"lzip2⋅⊥⋅ys = ⊥"
"lzip2⋅(lCons⋅x⋅xs)⋅⊥ = ⊥"
by fixrec_simp+
subsection ‹Mutual recursion with ‹fixrec››
text ‹Tree and forest types.›