src/CTT/ex/Synthesis.thy
author Manuel Eberl <eberlm@in.tum.de>
Tue, 16 Apr 2024 13:29:27 +0200
changeset 80107 247751d25102
parent 76539 8c94ca4dd035
permissions -rw-r--r--
canonical time function for List.nth
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     1
(*  Title:      CTT/ex/Synthesis.thy
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     3
    Copyright   1991  University of Cambridge
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     4
*)
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     5
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
     6
section \<open>Synthesis examples, using a crude form of narrowing\<close>
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     7
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
     8
theory Synthesis
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
     9
  imports "../CTT"
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    10
begin
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    11
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
    12
text \<open>discovery of predecessor function\<close>
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    13
schematic_goal "?a : \<Sum>pred:?A . Eq(N, pred`0, 0) \<times> (\<Prod>n:N. Eq(N, pred ` succ(n), n))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    14
  apply intr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    15
    apply eqintr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    16
    apply (rule_tac [3] reduction_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    17
      apply (rule_tac [5] comp_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    18
        apply rew
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    19
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    20
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
    21
text \<open>the function fst as an element of a function type\<close>
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 59498
diff changeset
    22
schematic_goal [folded basic_defs]:
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    23
  "A type \<Longrightarrow> ?a: \<Sum>f:?B . \<Prod>i:A. \<Prod>j:A. Eq(A, f ` <i,j>, i)"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    24
  apply intr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    25
   apply eqintr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    26
   apply (rule_tac [2] reduction_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    27
     apply (rule_tac [4] comp_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    28
       apply typechk
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    29
  txt "now put in A everywhere"
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    30
   apply assumption+
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    31
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    32
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
    33
text \<open>An interesting use of the eliminator, when\<close>
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    34
  (*The early implementation of unification caused non-rigid path in occur check
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    35
  See following example.*)
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    36
schematic_goal "?a : \<Prod>i:N. Eq(?A, ?b(inl(i)), <0    ,   i>)
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    37
                   \<times> Eq(?A, ?b(inr(i)), <succ(0), i>)"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    38
  apply intr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    39
   apply eqintr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    40
   apply (rule comp_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    41
     apply rew
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    42
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    43
19774
5fe7731d0836 allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents: 19761
diff changeset
    44
(*Here we allow the type to depend on i.
5fe7731d0836 allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents: 19761
diff changeset
    45
 This prevents the cycle in the first unification (no longer needed).
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    46
 Requires flex-flex to preserve the dependence.
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    47
 Simpler still: make ?A into a constant type N \<times> N.*)
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    48
schematic_goal "?a : \<Prod>i:N. Eq(?A(i), ?b(inl(i)), <0   ,   i>)
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    49
                  \<times>  Eq(?A(i), ?b(inr(i)), <succ(0),i>)"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    50
  oops
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    51
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
    52
  text \<open>A tricky combination of when and split\<close>
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    53
    (*Now handled easily, but caused great problems once*)
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 59498
diff changeset
    54
schematic_goal [folded basic_defs]:
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    55
  "?a : \<Prod>i:N. \<Prod>j:N. Eq(?A, ?b(inl(<i,j>)), i)
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    56
                           \<times>  Eq(?A, ?b(inr(<i,j>)), j)"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    57
  apply intr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    58
   apply eqintr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    59
   apply (rule PlusC_inl [THEN trans_elem])
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    60
      apply (rule_tac [4] comp_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    61
        apply (rule_tac [7] reduction_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    62
           apply (rule_tac [10] comp_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    63
             apply typechk
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    64
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    65
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    66
(*similar but allows the type to depend on i and j*)
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    67
schematic_goal "?a : \<Prod>i:N. \<Prod>j:N. Eq(?A(i,j), ?b(inl(<i,j>)), i)
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    68
                          \<times>   Eq(?A(i,j), ?b(inr(<i,j>)), j)"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    69
  oops
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    70
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    71
(*similar but specifying the type N simplifies the unification problems*)
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    72
schematic_goal "?a : \<Prod>i:N. \<Prod>j:N. Eq(N, ?b(inl(<i,j>)), i)
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    73
                          \<times>   Eq(N, ?b(inr(<i,j>)), j)"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    74
  oops
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    75
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    76
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
    77
  text \<open>Deriving the addition operator\<close>
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 59498
diff changeset
    78
schematic_goal [folded arith_defs]:
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    79
  "?c : \<Prod>n:N. Eq(N, ?f(0,n), n)
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    80
                  \<times>  (\<Prod>m:N. Eq(N, ?f(succ(m), n), succ(?f(m,n))))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    81
  apply intr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    82
   apply eqintr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    83
   apply (rule comp_rls)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    84
    apply rew
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    85
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
    86
76539
8c94ca4dd035 A new Isabelle/CTT example, and eliminated some old-style quotation marks
paulson <lp15@cam.ac.uk>
parents: 76377
diff changeset
    87
text \<open>The addition function -- using explicit lambdas\<close>
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 59498
diff changeset
    88
schematic_goal [folded arith_defs]:
61391
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    89
  "?c : \<Sum>plus : ?A .
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    90
         \<Prod>x:N. Eq(N, plus`0`x, x)
2332d9be352b tuned syntax -- more symbols;
wenzelm
parents: 61337
diff changeset
    91
                \<times>  (\<Prod>y:N. Eq(N, plus`succ(y)`x, succ(plus`y`x)))"
76377
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    92
  apply intr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    93
    apply eqintr
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    94
    apply (tactic "resolve_tac \<^context> [TSimp.split_eqn] 3")
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    95
     apply (tactic "SELECT_GOAL (rew_tac \<^context> []) 4")
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    96
         apply (tactic "resolve_tac \<^context> [TSimp.split_eqn] 3")
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    97
          apply (tactic "SELECT_GOAL (rew_tac \<^context> []) 4")
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    98
              apply (rule_tac [3] p = "y" in NC_succ)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    99
    (**  by (resolve_tac @{context} comp_rls 3);  caused excessive branching  **)
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   100
                apply rew
2510e6f7b11c Beautifying CTT a tiny bit
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   101
  done
19761
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   102
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   103
end
5cd82054c2c6 removed obsolete ML files;
wenzelm
parents:
diff changeset
   104