(* Title: ex/Synthesis.thy Author: Joshua Chen Date: 2018 Examples of synthesis *) theory Synthesis imports "../HoTT" begin section \Synthesis of the predecessor function\ text \ In this example we construct a predecessor function for the natural numbers. This is also done in @{file "~~/src/CTT/ex/Synthesis.thy"}, there the work is much easier as the equality type is extensional. \ text \First we show that the property we want is well-defined.\ lemma pred_welltyped: "\pred: \\\. (pred`0 =\<^sub>\ 0) \ (\n:\. pred`(succ n) =\<^sub>\ n): U O" by routine text \ Now we look for an inhabitant of this type. Observe that we're looking for a lambda term @{term pred} satisfying @{term "pred`0 =\<^sub>\ 0"} and @{term "\n:\. pred`(succ n) =\<^sub>\ n"}. What if we require *definitional* instead of just propositional equality? \ schematic_goal "?p`0 \ 0" and "\n. n: \ \ (?p`(succ n)) \ n" apply compute prefer 4 apply compute apply (rule Nat_routine | compute)+ oops \ \Something in the original proof broke when I revamped the theory. The completion of this derivation is left as an exercise to the reader.\ text \ The above proof finds a candidate, namely @{term "\n. ind\<^sub>\ (\a b. a) 0 n"}. We prove this has the required type and properties. \ definition pred :: t where "pred \ \<^bold>\n. ind\<^sub>\ (\a b. a) 0 n" lemma pred_type: "pred: \ \ \" unfolding pred_def by routine lemma pred_props: "\n. refl n>: (pred`0 =\<^sub>\ 0) \ (\n:\. pred`(succ n) =\<^sub>\ n)" unfolding pred_def by derive theorem "\n. refl(n)>>: \pred:\\\ . ((pred`0) =\<^sub>\ 0) \ (\n:\. (pred`(succ n)) =\<^sub>\ n)" by (derive lems: pred_type pred_props) end