(* Title: HoTT/ex/Synthesis.thy Author: Josh Chen Date: Aug 2018 Examples of synthesis. *) theory Synthesis imports "../HoTT" begin section \Synthesis of the predecessor function\ text " In this example we construct, with the help of Isabelle, a predecessor function for the natural numbers. This is also done in \CTT.thy\; there the work is easier as the equality type is extensional, and also the methods are set up a little more nicely. " 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 simple text " Now we look for an inhabitant of this type. Observe that we're looking for a lambda term \pred\ satisfying \(pred`0) =\<^sub>\ 0\ and \\n:\. (pred`(succ n)) =\<^sub>\ n\. What if we require **definitional** equality instead of just propositional equality? " schematic_goal "?p`0 \ 0" and "\n. n: \ \ (?p`(succ n)) \ n" apply compute prefer 4 apply compute prefer 3 apply (rule Nat_rules) apply (rule Nat_rules | assumption)+ done text " The above proof finds a candidate, namely \\<^bold>\n. ind\<^sub>\ (\a b. a) 0 n\. We prove this has the required type and properties. " definition pred :: Term where "pred \ \<^bold>\n. ind\<^sub>\ (\a b. a) 0 n" lemma pred_type: "pred: \ \ \" unfolding pred_def by simple lemma pred_props: "\n. refl(n)>: ((pred`0) =\<^sub>\ 0) \ (\n:\. (pred`(succ n)) =\<^sub>\ n)" proof (simple lems: pred_type) have *: "pred`0 \ 0" unfolding pred_def proof compute show "\n. n: \ \ ind\<^sub>\ (\a b. a) 0 n: \" by simple show "ind\<^sub>\ (\a b. a) 0 0 \ 0" proof (rule Nat_comps) show "\: U(O)" .. qed simple qed rule then show "refl(0): (pred`0) =\<^sub>\ 0" by (subst *) simple show "\<^bold>\n. refl(n): \n:\. (pred`(succ(n))) =\<^sub>\ n" unfolding pred_def proof show "\n. n: \ \ refl(n): ((\<^bold>\n. ind\<^sub>\ (\a b. a) 0 n)`succ(n)) =\<^sub>\ n" proof compute show "\n. n: \ \ ind\<^sub>\ (\a b. a) 0 n: \" by simple show "\n. n: \ \ refl(n): ind\<^sub>\ (\a b. a) 0 (succ n) =\<^sub>\ n" proof compute show "\: U(O)" .. qed simple qed rule qed rule qed theorem "\n. refl(n)>>: \pred:\\\ . ((pred`0) =\<^sub>\ 0) \ (\n:\. (pred`(succ n)) =\<^sub>\ n)" by (simple lems: pred_welltyped pred_type pred_props) end