# HG changeset patch # User paulson # Date 1669635535 0 # Node ID 8c94ca4dd035953eea71755ef59ab446b034bca1 # Parent 0bab4c7514785704b05cab3d4d54bd4f63e7f088 A new Isabelle/CTT example, and eliminated some old-style quotation marks diff -r 0bab4c751478 -r 8c94ca4dd035 src/CTT/ex/Elimination.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 \Examples with elimination rules\ theory Elimination imports "../CTT" begin -text "This finds the functions fst and snd!" +text \This finds the functions fst and snd!\ schematic_goal [folded basic_defs]: "A type \ ?a : (A \ A) \ A" apply pc @@ -23,7 +23,7 @@ back done -text "Double negation of the Excluded Middle" +text \Double negation of the Excluded Middle\ schematic_goal "A type \ ?a : ((A + (A\F)) \ F) \ F" apply intr apply (rule ProdE) @@ -31,13 +31,25 @@ apply pc done +text \Experiment: the proof above in Isar\ +lemma + assumes "A type" shows "(\<^bold>\f. f ` inr(\<^bold>\y. f ` inl(y))) : ((A + (A\F)) \ F) \ F" +proof intr + fix f + assume f: "f : A + (A \ F) \ F" + with assms have "inr(\<^bold>\y. f ` inl(y)) : A + (A \ F)" + by pc + then show "f ` inr(\<^bold>\y. f ` inl(y)) : F" + by (rule ProdE [OF f]) +qed (rule assms)+ + schematic_goal "\A type; B type\ \ ?a : (A \ B) \ (B \ A)" apply pc done (*The sequent version (ITT) could produce an interesting alternative by backtracking. No longer.*) -text "Binary sums and products" +text \Binary sums and products\ schematic_goal "\A type; B type; C type\ \ ?a : (A + B \ C) \ (A \ C) \ (B \ C)" apply pc done @@ -55,7 +67,7 @@ apply (pc assms) done -text "Construction of the currying functional" +text \Construction of the currying functional\ schematic_goal "\A type; B type; C type\ \ ?a : (A \ B \ C) \ (A \ (B \ 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 \Martin-Löf (1984), page 48: axiom of sum-elimination (uncurry)\ schematic_goal "\A type; B type; C type\ \ ?a : (A \ (B \ C)) \ (A \ B \ C)" apply pc done @@ -85,12 +97,12 @@ apply (pc assms) done -text "Function application" +text \Function application\ schematic_goal "\A type; B type\ \ ?a : ((A \ B) \ A) \ B" apply pc done -text "Basic test of quantifier reasoning" +text \Basic test of quantifier reasoning\ 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 \Martin-Löf (1984) pages 36-7: the combinator S\ schematic_goal assumes "A type" and "\x. x:A \ B(x) type" @@ -111,7 +123,7 @@ apply (pc assms) done -text "Martin-Löf (1984) page 58: the axiom of disjunction elimination" +text \Martin-Löf (1984) page 58: the axiom of disjunction elimination\ 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 \AXIOM OF CHOICE! Delicate use of elimination rules\ schematic_goal assumes "A type" and "\x. x:A \ B(x) type" @@ -168,7 +180,7 @@ by (intro replace_type [OF subst_eqtyparg]) (typechk SumE_fst assms \a : A\) qed -text "Axiom of choice. Proof without fst, snd. Harder still!" +text \Axiom of choice. Proof without fst, snd. Harder still!\ schematic_goal [folded basic_defs]: assumes "A type" and "\x. x:A \ B(x) type" @@ -192,7 +204,7 @@ apply assumption done -text "Example of sequent-style deduction" +text \Example of sequent-style deduction\ (*When splitting z:A \ B, the assumption C(z) is affected; ?a becomes \<^bold>\u. split(u,\v w.split(v,\x y.\<^bold> \z. >) ` w) *) schematic_goal diff -r 0bab4c751478 -r 8c94ca4dd035 src/CTT/ex/Synthesis.thy --- 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 \Synthesis examples, using a crude form of narrowing\ theory Synthesis imports "../CTT" begin -text "discovery of predecessor function" +text \discovery of predecessor function\ schematic_goal "?a : \pred:?A . Eq(N, pred`0, 0) \ (\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 \the function fst as an element of a function type\ schematic_goal [folded basic_defs]: "A type \ ?a: \f:?B . \i:A. \j:A. Eq(A, f ` , i)" apply intr @@ -30,7 +30,7 @@ apply assumption+ done -text "An interesting use of the eliminator, when" +text \An interesting use of the eliminator, when\ (*The early implementation of unification caused non-rigid path in occur check See following example.*) schematic_goal "?a : \i:N. Eq(?A, ?b(inl(i)), <0 , i>) @@ -49,7 +49,7 @@ \ Eq(?A(i), ?b(inr(i)), )" oops - text "A tricky combination of when and split" + text \A tricky combination of when and split\ (*Now handled easily, but caused great problems once*) schematic_goal [folded basic_defs]: "?a : \i:N. \j:N. Eq(?A, ?b(inl()), i) @@ -74,7 +74,7 @@ oops - text "Deriving the addition operator" + text \Deriving the addition operator\ schematic_goal [folded arith_defs]: "?c : \n:N. Eq(N, ?f(0,n), n) \ (\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 \The addition function -- using explicit lambdas\ schematic_goal [folded arith_defs]: "?c : \plus : ?A . \x:N. Eq(N, plus`0`x, x) diff -r 0bab4c751478 -r 8c94ca4dd035 src/CTT/ex/Typechecking.thy --- 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 \Easy examples: type checking and type deduction\ theory Typechecking imports "../CTT" @@ -44,15 +44,15 @@ schematic_goal "\x:N . \y:N . Eq(?A,x,y) type" apply typechk done -text "typechecking an application of fst" +text \typechecking an application of fst\ schematic_goal "(\<^bold>\u. split(u, \v w. v)) ` <0, succ(0)> : ?A" apply typechk done -text "typechecking the predecessor function" +text \typechecking the predecessor function\ schematic_goal "\<^bold>\n. rec(n, 0, \x y. x) : ?A" apply typechk done -text "typechecking the addition function" +text \typechecking the addition function\ schematic_goal "\<^bold>\n. \<^bold>\m. rec(n, m, \x y. succ(y)) : ?A" apply typechk done @@ -71,7 +71,7 @@ apply N done -text "typechecking fst (as a function object)" +text \typechecking fst (as a function object)\ schematic_goal "\<^bold>\i. split(i, \j k. j) : ?A" apply typechk apply N