From e6473c383b479610cee4c0119e5811a12a938cf9 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Tue, 14 Aug 2018 17:43:03 +0200 Subject: Well-formation rules are back in the methods; new theory synthesizing the natural number predecessor function. --- ex/Synthesis.thy | 81 ++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 81 insertions(+) create mode 100644 ex/Synthesis.thy (limited to 'ex/Synthesis.thy') diff --git a/ex/Synthesis.thy b/ex/Synthesis.thy new file mode 100644 index 0000000..60655e5 --- /dev/null +++ b/ex/Synthesis.thy @@ -0,0 +1,81 @@ +(* 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 (subst comp, rule Nat_rules) +prefer 3 apply (subst comp, rule Nat_rules) +prefer 3 apply (rule Nat_rules) +prefer 8 apply (rule Nat_rules | assumption)+ +done + +text " + The above proof finds a candidate, namely \\<^bold>\n. ind\<^sub>\ (\a b. a) n n\. + We prove this has the required type and properties. +" + +definition pred :: Term where "pred \ \<^bold>\n. ind\<^sub>\ (\a b. a) n 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 lem: pred_type) + have *: "pred`0 \ 0" unfolding pred_def + proof (subst comp) + show "\n. n: \ \ ind\<^sub>\ (\a b. a) n 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) n n)`succ(n)) =\<^sub>\ n" + proof (subst comp) + show "\n. n: \ \ ind\<^sub>\ (\a b. a) n n: \" by simple + show "\n. n: \ \ refl(n): ind\<^sub>\ (\a b. a) (succ n) (succ n) =\<^sub>\ n" + proof (subst comp) + 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 lem: pred_welltyped pred_type pred_props) + + +end \ No newline at end of file -- cgit v1.2.3