A new Isabelle/CTT example, and eliminated some old-style quotation marks
authorpaulson <lp15@cam.ac.uk>
Mon, 28 Nov 2022 11:38:55 +0000
changeset 76539 8c94ca4dd035
parent 76538 0bab4c751478
child 76540 83de6e9ae983
A new Isabelle/CTT example, and eliminated some old-style quotation marks
src/CTT/ex/Elimination.thy
src/CTT/ex/Synthesis.thy
src/CTT/ex/Typechecking.thy
--- a/src/CTT/ex/Elimination.thy	Fri Nov 25 22:38:10 2022 +0100
+++ b/src/CTT/ex/Elimination.thy	Mon Nov 28 11:38:55 2022 +0000
@@ -6,13 +6,13 @@
 (Bibliopolis, 1984).
 *)
 
-section "Examples with elimination rules"
+section \<open>Examples with elimination rules\<close>
 
 theory Elimination
 imports "../CTT"
 begin
 
-text "This finds the functions fst and snd!"
+text \<open>This finds the functions fst and snd!\<close>
 
 schematic_goal [folded basic_defs]: "A type \<Longrightarrow> ?a : (A \<times> A) \<longrightarrow> A"
 apply pc
@@ -23,7 +23,7 @@
   back
   done
 
-text "Double negation of the Excluded Middle"
+text \<open>Double negation of the Excluded Middle\<close>
 schematic_goal "A type \<Longrightarrow> ?a : ((A + (A\<longrightarrow>F)) \<longrightarrow> F) \<longrightarrow> F"
   apply intr
   apply (rule ProdE)
@@ -31,13 +31,25 @@
   apply pc
   done
 
+text \<open>Experiment: the proof above in Isar\<close>
+lemma
+  assumes "A type" shows "(\<^bold>\<lambda>f. f ` inr(\<^bold>\<lambda>y. f ` inl(y))) : ((A + (A\<longrightarrow>F)) \<longrightarrow> F) \<longrightarrow> F"
+proof intr
+  fix f
+  assume f: "f : A + (A \<longrightarrow> F) \<longrightarrow> F" 
+  with assms have "inr(\<^bold>\<lambda>y. f ` inl(y)) : A + (A \<longrightarrow> F)"
+    by pc
+  then show "f ` inr(\<^bold>\<lambda>y. f ` inl(y)) : F" 
+    by (rule ProdE [OF f])
+qed (rule assms)+
+
 schematic_goal "\<lbrakk>A type; B type\<rbrakk> \<Longrightarrow> ?a : (A \<times> B) \<longrightarrow> (B \<times> A)"
 apply pc
 done
 (*The sequent version (ITT) could produce an interesting alternative
   by backtracking.  No longer.*)
 
-text "Binary sums and products"
+text \<open>Binary sums and products\<close>
 schematic_goal "\<lbrakk>A type; B type; C type\<rbrakk> \<Longrightarrow> ?a : (A + B \<longrightarrow> C) \<longrightarrow> (A \<longrightarrow> C) \<times> (B \<longrightarrow> C)"
   apply pc
   done
@@ -55,7 +67,7 @@
   apply (pc assms)
   done
 
-text "Construction of the currying functional"
+text \<open>Construction of the currying functional\<close>
 schematic_goal "\<lbrakk>A type; B type; C type\<rbrakk> \<Longrightarrow> ?a : (A \<times> B \<longrightarrow> C) \<longrightarrow> (A \<longrightarrow> (B \<longrightarrow> C))"
   apply pc
   done
@@ -70,7 +82,7 @@
   apply (pc assms)
   done
 
-text "Martin-Löf (1984), page 48: axiom of sum-elimination (uncurry)"
+text \<open>Martin-Löf (1984), page 48: axiom of sum-elimination (uncurry)\<close>
 schematic_goal "\<lbrakk>A type; B type; C type\<rbrakk> \<Longrightarrow> ?a : (A \<longrightarrow> (B \<longrightarrow> C)) \<longrightarrow> (A \<times> B \<longrightarrow> C)"
   apply pc
   done
@@ -85,12 +97,12 @@
   apply (pc assms)
   done
 
-text "Function application"
+text \<open>Function application\<close>
 schematic_goal "\<lbrakk>A type; B type\<rbrakk> \<Longrightarrow> ?a : ((A \<longrightarrow> B) \<times> A) \<longrightarrow> B"
   apply pc
   done
 
-text "Basic test of quantifier reasoning"
+text \<open>Basic test of quantifier reasoning\<close>
 schematic_goal
   assumes "A type"
     and "B type"
@@ -101,7 +113,7 @@
   apply (pc assms)
   done
 
-text "Martin-Löf (1984) pages 36-7: the combinator S"
+text \<open>Martin-Löf (1984) pages 36-7: the combinator S\<close>
 schematic_goal
   assumes "A type"
     and "\<And>x. x:A \<Longrightarrow> B(x) type"
@@ -111,7 +123,7 @@
   apply (pc assms)
   done
 
-text "Martin-Löf (1984) page 58: the axiom of disjunction elimination"
+text \<open>Martin-Löf (1984) page 58: the axiom of disjunction elimination\<close>
 schematic_goal
   assumes "A type"
     and "B type"
@@ -129,7 +141,7 @@
 
 
 (*Martin-Löf (1984) page 50*)
-text "AXIOM OF CHOICE!  Delicate use of elimination rules"
+text \<open>AXIOM OF CHOICE!  Delicate use of elimination rules\<close>
 schematic_goal
   assumes "A type"
     and "\<And>x. x:A \<Longrightarrow> B(x) type"
@@ -168,7 +180,7 @@
     by (intro replace_type [OF subst_eqtyparg]) (typechk SumE_fst assms \<open>a : A\<close>)
 qed
 
-text "Axiom of choice.  Proof without fst, snd.  Harder still!"
+text \<open>Axiom of choice.  Proof without fst, snd.  Harder still!\<close>
 schematic_goal [folded basic_defs]:
   assumes "A type"
     and "\<And>x. x:A \<Longrightarrow> B(x) type"
@@ -192,7 +204,7 @@
   apply assumption
   done
 
-text "Example of sequent-style deduction"
+text \<open>Example of sequent-style deduction\<close>
   (*When splitting z:A \<times> B, the assumption C(z) is affected;  ?a becomes
     \<^bold>\<lambda>u. split(u,\<lambda>v w.split(v,\<lambda>x y.\<^bold> \<lambda>z. <x,<y,z>>) ` w)     *)
 schematic_goal
--- a/src/CTT/ex/Synthesis.thy	Fri Nov 25 22:38:10 2022 +0100
+++ b/src/CTT/ex/Synthesis.thy	Mon Nov 28 11:38:55 2022 +0000
@@ -3,13 +3,13 @@
     Copyright   1991  University of Cambridge
 *)
 
-section "Synthesis examples, using a crude form of narrowing"
+section \<open>Synthesis examples, using a crude form of narrowing\<close>
 
 theory Synthesis
   imports "../CTT"
 begin
 
-text "discovery of predecessor function"
+text \<open>discovery of predecessor function\<close>
 schematic_goal "?a : \<Sum>pred:?A . Eq(N, pred`0, 0) \<times> (\<Prod>n:N. Eq(N, pred ` succ(n), n))"
   apply intr
     apply eqintr
@@ -18,7 +18,7 @@
         apply rew
   done
 
-text "the function fst as an element of a function type"
+text \<open>the function fst as an element of a function type\<close>
 schematic_goal [folded basic_defs]:
   "A type \<Longrightarrow> ?a: \<Sum>f:?B . \<Prod>i:A. \<Prod>j:A. Eq(A, f ` <i,j>, i)"
   apply intr
@@ -30,7 +30,7 @@
    apply assumption+
   done
 
-text "An interesting use of the eliminator, when"
+text \<open>An interesting use of the eliminator, when\<close>
   (*The early implementation of unification caused non-rigid path in occur check
   See following example.*)
 schematic_goal "?a : \<Prod>i:N. Eq(?A, ?b(inl(i)), <0    ,   i>)
@@ -49,7 +49,7 @@
                   \<times>  Eq(?A(i), ?b(inr(i)), <succ(0),i>)"
   oops
 
-  text "A tricky combination of when and split"
+  text \<open>A tricky combination of when and split\<close>
     (*Now handled easily, but caused great problems once*)
 schematic_goal [folded basic_defs]:
   "?a : \<Prod>i:N. \<Prod>j:N. Eq(?A, ?b(inl(<i,j>)), i)
@@ -74,7 +74,7 @@
   oops
 
 
-  text "Deriving the addition operator"
+  text \<open>Deriving the addition operator\<close>
 schematic_goal [folded arith_defs]:
   "?c : \<Prod>n:N. Eq(N, ?f(0,n), n)
                   \<times>  (\<Prod>m:N. Eq(N, ?f(succ(m), n), succ(?f(m,n))))"
@@ -84,7 +84,7 @@
     apply rew
   done
 
-text "The addition function -- using explicit lambdas"
+text \<open>The addition function -- using explicit lambdas\<close>
 schematic_goal [folded arith_defs]:
   "?c : \<Sum>plus : ?A .
          \<Prod>x:N. Eq(N, plus`0`x, x)
--- a/src/CTT/ex/Typechecking.thy	Fri Nov 25 22:38:10 2022 +0100
+++ b/src/CTT/ex/Typechecking.thy	Mon Nov 28 11:38:55 2022 +0000
@@ -3,7 +3,7 @@
     Copyright   1991  University of Cambridge
 *)
 
-section "Easy examples: type checking and type deduction"
+section \<open>Easy examples: type checking and type deduction\<close>
 
 theory Typechecking
 imports "../CTT"
@@ -44,15 +44,15 @@
 schematic_goal "\<Prod>x:N . \<Prod>y:N . Eq(?A,x,y) type"
   apply typechk done
 
-text "typechecking an application of fst"
+text \<open>typechecking an application of fst\<close>
 schematic_goal "(\<^bold>\<lambda>u. split(u, \<lambda>v w. v)) ` <0, succ(0)> : ?A"
   apply typechk done
 
-text "typechecking the predecessor function"
+text \<open>typechecking the predecessor function\<close>
 schematic_goal "\<^bold>\<lambda>n. rec(n, 0, \<lambda>x y. x) : ?A"
   apply typechk done
 
-text "typechecking the addition function"
+text \<open>typechecking the addition function\<close>
 schematic_goal "\<^bold>\<lambda>n. \<^bold>\<lambda>m. rec(n, m, \<lambda>x y. succ(y)) : ?A"
   apply typechk done
 
@@ -71,7 +71,7 @@
    apply N
   done
 
-text "typechecking fst (as a function object)"
+text \<open>typechecking fst (as a function object)\<close>
 schematic_goal "\<^bold>\<lambda>i. split(i, \<lambda>j k. j) : ?A"
   apply typechk 
    apply N