| author | haftmann | 
| Tue, 31 Aug 2010 13:15:35 +0200 | |
| changeset 38922 | ec2a8efd8990 | 
| parent 36319 | 8feb2c4bef1a | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 19761 | 1 | (* Title: CTT/ex/Synthesis.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1991 University of Cambridge | |
| 4 | *) | |
| 5 | ||
| 6 | header "Synthesis examples, using a crude form of narrowing" | |
| 7 | ||
| 8 | theory Synthesis | |
| 9 | imports Arith | |
| 10 | begin | |
| 11 | ||
| 12 | text "discovery of predecessor function" | |
| 36319 | 13 | schematic_lemma "?a : SUM pred:?A . Eq(N, pred`0, 0) | 
| 19761 | 14 | * (PROD n:N. Eq(N, pred ` succ(n), n))" | 
| 15 | apply (tactic "intr_tac []") | |
| 16 | apply (tactic eqintr_tac) | |
| 17 | apply (rule_tac [3] reduction_rls) | |
| 18 | apply (rule_tac [5] comp_rls) | |
| 19 | apply (tactic "rew_tac []") | |
| 20 | done | |
| 21 | ||
| 22 | text "the function fst as an element of a function type" | |
| 36319 | 23 | schematic_lemma [folded basic_defs]: | 
| 19761 | 24 | "A type ==> ?a: SUM f:?B . PROD i:A. PROD j:A. Eq(A, f ` <i,j>, i)" | 
| 25 | apply (tactic "intr_tac []") | |
| 26 | apply (tactic eqintr_tac) | |
| 27 | apply (rule_tac [2] reduction_rls) | |
| 28 | apply (rule_tac [4] comp_rls) | |
| 29 | apply (tactic "typechk_tac []") | |
| 30 | txt "now put in A everywhere" | |
| 31 | apply assumption+ | |
| 32 | done | |
| 33 | ||
| 34 | text "An interesting use of the eliminator, when" | |
| 35 | (*The early implementation of unification caused non-rigid path in occur check | |
| 36 | See following example.*) | |
| 36319 | 37 | schematic_lemma "?a : PROD i:N. Eq(?A, ?b(inl(i)), <0 , i>) | 
| 19761 | 38 | * Eq(?A, ?b(inr(i)), <succ(0), i>)" | 
| 39 | apply (tactic "intr_tac []") | |
| 40 | apply (tactic eqintr_tac) | |
| 41 | apply (rule comp_rls) | |
| 42 | apply (tactic "rew_tac []") | |
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19761diff
changeset | 43 | done | 
| 19761 | 44 | |
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19761diff
changeset | 45 | (*Here we allow the type to depend on i. | 
| 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19761diff
changeset | 46 | This prevents the cycle in the first unification (no longer needed). | 
| 19761 | 47 | Requires flex-flex to preserve the dependence. | 
| 48 | Simpler still: make ?A into a constant type N*N.*) | |
| 36319 | 49 | schematic_lemma "?a : PROD i:N. Eq(?A(i), ?b(inl(i)), <0 , i>) | 
| 19761 | 50 | * Eq(?A(i), ?b(inr(i)), <succ(0),i>)" | 
| 51 | oops | |
| 52 | ||
| 53 | text "A tricky combination of when and split" | |
| 54 | (*Now handled easily, but caused great problems once*) | |
| 36319 | 55 | schematic_lemma [folded basic_defs]: | 
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19761diff
changeset | 56 | "?a : PROD i:N. PROD j:N. Eq(?A, ?b(inl(<i,j>)), i) | 
| 19761 | 57 | * Eq(?A, ?b(inr(<i,j>)), j)" | 
| 58 | apply (tactic "intr_tac []") | |
| 59 | apply (tactic eqintr_tac) | |
| 60 | apply (rule PlusC_inl [THEN trans_elem]) | |
| 61 | apply (rule_tac [4] comp_rls) | |
| 62 | apply (rule_tac [7] reduction_rls) | |
| 63 | apply (rule_tac [10] comp_rls) | |
| 64 | apply (tactic "typechk_tac []") | |
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19761diff
changeset | 65 | done | 
| 19761 | 66 | |
| 67 | (*similar but allows the type to depend on i and j*) | |
| 36319 | 68 | schematic_lemma "?a : PROD i:N. PROD j:N. Eq(?A(i,j), ?b(inl(<i,j>)), i) | 
| 19761 | 69 | * Eq(?A(i,j), ?b(inr(<i,j>)), j)" | 
| 70 | oops | |
| 71 | ||
| 72 | (*similar but specifying the type N simplifies the unification problems*) | |
| 36319 | 73 | schematic_lemma "?a : PROD i:N. PROD j:N. Eq(N, ?b(inl(<i,j>)), i) | 
| 19761 | 74 | * Eq(N, ?b(inr(<i,j>)), j)" | 
| 75 | oops | |
| 76 | ||
| 77 | ||
| 78 | text "Deriving the addition operator" | |
| 36319 | 79 | schematic_lemma [folded arith_defs]: | 
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19761diff
changeset | 80 | "?c : PROD n:N. Eq(N, ?f(0,n), n) | 
| 19761 | 81 | * (PROD m:N. Eq(N, ?f(succ(m), n), succ(?f(m,n))))" | 
| 82 | apply (tactic "intr_tac []") | |
| 83 | apply (tactic eqintr_tac) | |
| 84 | apply (rule comp_rls) | |
| 85 | apply (tactic "rew_tac []") | |
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19761diff
changeset | 86 | done | 
| 19761 | 87 | |
| 88 | text "The addition function -- using explicit lambdas" | |
| 36319 | 89 | schematic_lemma [folded arith_defs]: | 
| 19774 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19761diff
changeset | 90 | "?c : SUM plus : ?A . | 
| 
5fe7731d0836
allow non-trivial schematic goals (via embedded term vars);
 wenzelm parents: 
19761diff
changeset | 91 | PROD x:N. Eq(N, plus`0`x, x) | 
| 19761 | 92 | * (PROD y:N. Eq(N, plus`succ(y)`x, succ(plus`y`x)))" | 
| 93 | apply (tactic "intr_tac []") | |
| 94 | apply (tactic eqintr_tac) | |
| 95 | apply (tactic "resolve_tac [TSimp.split_eqn] 3") | |
| 96 | apply (tactic "SELECT_GOAL (rew_tac []) 4") | |
| 97 | apply (tactic "resolve_tac [TSimp.split_eqn] 3") | |
| 98 | apply (tactic "SELECT_GOAL (rew_tac []) 4") | |
| 99 | apply (rule_tac [3] p = "y" in NC_succ) | |
| 100 | (** by (resolve_tac comp_rls 3); caused excessive branching **) | |
| 101 | apply (tactic "rew_tac []") | |
| 102 | done | |
| 103 | ||
| 104 | end | |
| 105 |