src/CTT/ex/Synthesis.thy
changeset 61391 2332d9be352b
parent 61337 4645502c3c64
child 65447 fae6051ec192
--- a/src/CTT/ex/Synthesis.thy	Sat Oct 10 21:14:00 2015 +0200
+++ b/src/CTT/ex/Synthesis.thy	Sat Oct 10 21:43:07 2015 +0200
@@ -10,8 +10,7 @@
 begin
 
 text "discovery of predecessor function"
-schematic_goal "?a : SUM pred:?A . Eq(N, pred`0, 0)
-                  *  (PROD n:N. Eq(N, pred ` succ(n), n))"
+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
 apply (rule_tac [3] reduction_rls)
@@ -21,7 +20,7 @@
 
 text "the function fst as an element of a function type"
 schematic_goal [folded basic_defs]:
-  "A type \<Longrightarrow> ?a: SUM f:?B . PROD i:A. PROD j:A. Eq(A, f ` <i,j>, i)"
+  "A type \<Longrightarrow> ?a: \<Sum>f:?B . \<Prod>i:A. \<Prod>j:A. Eq(A, f ` <i,j>, i)"
 apply intr
 apply eqintr
 apply (rule_tac [2] reduction_rls)
@@ -34,8 +33,8 @@
 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 : PROD i:N. Eq(?A, ?b(inl(i)), <0    ,   i>)
-                   * Eq(?A, ?b(inr(i)), <succ(0), i>)"
+schematic_goal "?a : \<Prod>i:N. Eq(?A, ?b(inl(i)), <0    ,   i>)
+                   \<times> Eq(?A, ?b(inr(i)), <succ(0), i>)"
 apply intr
 apply eqintr
 apply (rule comp_rls)
@@ -45,16 +44,16 @@
 (*Here we allow the type to depend on i.
  This prevents the cycle in the first unification (no longer needed).
  Requires flex-flex to preserve the dependence.
- Simpler still: make ?A into a constant type N*N.*)
-schematic_goal "?a : PROD i:N. Eq(?A(i), ?b(inl(i)), <0   ,   i>)
-                  *  Eq(?A(i), ?b(inr(i)), <succ(0),i>)"
+ Simpler still: make ?A into a constant type N \<times> N.*)
+schematic_goal "?a : \<Prod>i:N. Eq(?A(i), ?b(inl(i)), <0   ,   i>)
+                  \<times>  Eq(?A(i), ?b(inr(i)), <succ(0),i>)"
 oops
 
 text "A tricky combination of when and split"
 (*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)
-                           *  Eq(?A, ?b(inr(<i,j>)), j)"
+  "?a : \<Prod>i:N. \<Prod>j:N. Eq(?A, ?b(inl(<i,j>)), i)
+                           \<times>  Eq(?A, ?b(inr(<i,j>)), j)"
 apply intr
 apply eqintr
 apply (rule PlusC_inl [THEN trans_elem])
@@ -65,20 +64,20 @@
 done
 
 (*similar but allows the type to depend on i and j*)
-schematic_goal "?a : PROD i:N. PROD j:N. Eq(?A(i,j), ?b(inl(<i,j>)), i)
-                          *   Eq(?A(i,j), ?b(inr(<i,j>)), j)"
+schematic_goal "?a : \<Prod>i:N. \<Prod>j:N. Eq(?A(i,j), ?b(inl(<i,j>)), i)
+                          \<times>   Eq(?A(i,j), ?b(inr(<i,j>)), j)"
 oops
 
 (*similar but specifying the type N simplifies the unification problems*)
-schematic_goal "?a : PROD i:N. PROD j:N. Eq(N, ?b(inl(<i,j>)), i)
-                          *   Eq(N, ?b(inr(<i,j>)), j)"
+schematic_goal "?a : \<Prod>i:N. \<Prod>j:N. Eq(N, ?b(inl(<i,j>)), i)
+                          \<times>   Eq(N, ?b(inr(<i,j>)), j)"
 oops
 
 
 text "Deriving the addition operator"
 schematic_goal [folded arith_defs]:
-  "?c : PROD n:N. Eq(N, ?f(0,n), n)
-                  *  (PROD m:N. Eq(N, ?f(succ(m), n), succ(?f(m,n))))"
+  "?c : \<Prod>n:N. Eq(N, ?f(0,n), n)
+                  \<times>  (\<Prod>m:N. Eq(N, ?f(succ(m), n), succ(?f(m,n))))"
 apply intr
 apply eqintr
 apply (rule comp_rls)
@@ -87,9 +86,9 @@
 
 text "The addition function -- using explicit lambdas"
 schematic_goal [folded arith_defs]:
-  "?c : SUM plus : ?A .
-         PROD x:N. Eq(N, plus`0`x, x)
-                *  (PROD y:N. Eq(N, plus`succ(y)`x, succ(plus`y`x)))"
+  "?c : \<Sum>plus : ?A .
+         \<Prod>x:N. Eq(N, plus`0`x, x)
+                \<times>  (\<Prod>y:N. Eq(N, plus`succ(y)`x, succ(plus`y`x)))"
 apply intr
 apply eqintr
 apply (tactic "resolve_tac @{context} [TSimp.split_eqn] 3")