| author | wenzelm | 
| Sun, 28 May 2000 21:55:50 +0200 | |
| changeset 8991 | dc70b797827f | 
| parent 6141 | a6922171b396 | 
| child 10179 | 9d5678e6bf34 | 
| permissions | -rw-r--r-- | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 1 | (* Title: HOL/ex/comb.ML | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 4 | Copyright 1996 University of Cambridge | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 5 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 6 | Combinatory Logic example: the Church-Rosser Theorem | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 7 | Curiously, combinators do not include free variables. | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 8 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 9 | Example taken from | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 10 | J. Camilleri and T. F. Melham. | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 11 | Reasoning with Inductively Defined Relations in the HOL Theorem Prover. | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 12 | Report 265, University of Cambridge Computer Laboratory, 1992. | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 13 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 14 | HOL system proofs may be found in | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 15 | /usr/groups/theory/hvg-aftp/contrib/rule-induction/cl.ml | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 16 | *) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 17 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 18 | (*** Reflexive/Transitive closure preserves the Church-Rosser property | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 19 | So does the Transitive closure; use r_into_trancl instead of rtrancl_refl | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 20 | ***) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 21 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 22 | val [_, spec_mp] = [spec] RL [mp]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 23 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 24 | (*Strip lemma. The induction hyp is all but the last diamond of the strip.*) | 
| 5069 | 25 | Goalw [diamond_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 26 | "[| diamond(r); (x,y):r^* |] ==> \ | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 27 | \ ALL y'. (x,y'):r --> (EX z. (y',z): r^* & (y,z): r)"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 28 | by (etac rtrancl_induct 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 29 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 30 | by (slow_best_tac (set_cs addIs [r_into_rtrancl RSN (2, rtrancl_trans)] | 
| 4301 | 31 | addSDs [spec_mp]) 1); | 
| 5360 | 32 | qed_spec_mp "diamond_strip_lemmaE"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 33 | |
| 5360 | 34 | Goal "diamond(r) ==> diamond(r^*)"; | 
| 35 | by (stac diamond_def 1); (*unfold only in goal, not in premise!*) | |
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 36 | by (rtac (impI RS allI RS allI) 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 37 | by (etac rtrancl_induct 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 38 | by (Blast_tac 1); | 
| 5360 | 39 | by (best_tac (*Seems to be a brittle, undirected search*) | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 40 | (set_cs addIs [r_into_rtrancl RSN (2, rtrancl_trans)] | 
| 5360 | 41 | addEs [diamond_strip_lemmaE RS exE]) 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 42 | qed "diamond_rtrancl"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 43 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 44 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 45 | (*** Results about Contraction ***) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 46 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 47 | (*Derive a case for each combinator constructor*) | 
| 6141 | 48 | val K_contractE = contract.mk_cases "K -1-> z"; | 
| 49 | val S_contractE = contract.mk_cases "S -1-> z"; | |
| 50 | val Ap_contractE = contract.mk_cases "x#y -1-> z"; | |
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 51 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 52 | AddSIs [contract.K, contract.S]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 53 | AddIs [contract.Ap1, contract.Ap2]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 54 | AddSEs [K_contractE, S_contractE, Ap_contractE]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 55 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 56 | Goalw [I_def] "I -1-> z ==> P"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 57 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 58 | qed "I_contract_E"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 59 | AddSEs [I_contract_E]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 60 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 61 | Goal "K#x -1-> z ==> (EX x'. z = K#x' & x -1-> x')"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 62 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 63 | qed "K1_contractD"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 64 | AddSEs [K1_contractD]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 65 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 66 | Goal "x ---> y ==> x#z ---> y#z"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 67 | by (etac rtrancl_induct 1); | 
| 4089 | 68 | by (ALLGOALS (blast_tac (claset() addIs [r_into_rtrancl, rtrancl_trans]))); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 69 | qed "Ap_reduce1"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 70 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 71 | Goal "x ---> y ==> z#x ---> z#y"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 72 | by (etac rtrancl_induct 1); | 
| 4089 | 73 | by (ALLGOALS (blast_tac (claset() addIs [r_into_rtrancl, rtrancl_trans]))); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 74 | qed "Ap_reduce2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 75 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 76 | (** Counterexample to the diamond property for -1-> **) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 77 | |
| 5069 | 78 | Goal "K#I#(I#I) -1-> I"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 79 | by (rtac contract.K 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 80 | qed "KIII_contract1"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 81 | |
| 5069 | 82 | Goalw [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 83 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 84 | qed "KIII_contract2"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 85 | |
| 5069 | 86 | Goal "K#I#((K#I)#(K#I)) -1-> I"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 87 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 88 | qed "KIII_contract3"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 89 | |
| 5069 | 90 | Goalw [diamond_def] "~ diamond(contract)"; | 
| 4089 | 91 | by (blast_tac (claset() addIs [KIII_contract1,KIII_contract2,KIII_contract3]) 1); | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 92 | qed "not_diamond_contract"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 93 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 94 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 95 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 96 | (*** Results about Parallel Contraction ***) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 97 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 98 | (*Derive a case for each combinator constructor*) | 
| 6141 | 99 | val K_parcontractE = parcontract.mk_cases "K =1=> z"; | 
| 100 | val S_parcontractE = parcontract.mk_cases "S =1=> z"; | |
| 101 | val Ap_parcontractE = parcontract.mk_cases "x#y =1=> z"; | |
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 102 | |
| 4301 | 103 | AddSIs [parcontract.refl, parcontract.K, parcontract.S]; | 
| 104 | AddIs [parcontract.Ap]; | |
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 105 | AddSEs [K_parcontractE, S_parcontractE,Ap_parcontractE]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 106 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 107 | (*** Basic properties of parallel contraction ***) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 108 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 109 | Goal "K#x =1=> z ==> (EX x'. z = K#x' & x =1=> x')"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 110 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 111 | qed "K1_parcontractD"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 112 | AddSDs [K1_parcontractD]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 113 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 114 | Goal "S#x =1=> z ==> (EX x'. z = S#x' & x =1=> x')"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 115 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 116 | qed "S1_parcontractD"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 117 | AddSDs [S1_parcontractD]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 118 | |
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 119 | Goal "S#x#y =1=> z ==> (EX x' y'. z = S#x'#y' & x =1=> x' & y =1=> y')"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 120 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 121 | qed "S2_parcontractD"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 122 | AddSDs [S2_parcontractD]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 123 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 124 | (*The rules above are not essential but make proofs much faster*) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 125 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 126 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 127 | (*Church-Rosser property for parallel contraction*) | 
| 5069 | 128 | Goalw [diamond_def] "diamond parcontract"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 129 | by (rtac (impI RS allI RS allI) 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 130 | by (etac parcontract.induct 1 THEN prune_params_tac); | 
| 3724 | 131 | by Safe_tac; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 132 | by (ALLGOALS Blast_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 133 | qed "diamond_parcontract"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 134 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 135 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 136 | (*** Equivalence of x--->y and x===>y ***) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 137 | |
| 5069 | 138 | Goal "contract <= parcontract"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 139 | by (rtac subsetI 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 140 | by (split_all_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 141 | by (etac contract.induct 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 142 | by (ALLGOALS Blast_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 143 | qed "contract_subset_parcontract"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 144 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 145 | (*Reductions: simply throw together reflexivity, transitivity and | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 146 | the one-step reductions*) | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 147 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 148 | AddIs [Ap_reduce1, Ap_reduce2, r_into_rtrancl, rtrancl_trans]; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 149 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 150 | (*Example only: not used*) | 
| 5069 | 151 | Goalw [I_def] "I#x ---> x"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 152 | by (Blast_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 153 | qed "reduce_I"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 154 | |
| 5069 | 155 | Goal "parcontract <= contract^*"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 156 | by (rtac subsetI 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 157 | by (split_all_tac 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 158 | by (etac parcontract.induct 1 THEN prune_params_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 159 | by (ALLGOALS Blast_tac); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 160 | qed "parcontract_subset_reduce"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 161 | |
| 5069 | 162 | Goal "contract^* = parcontract^*"; | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 163 | by (REPEAT | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 164 | (resolve_tac [equalityI, | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 165 | contract_subset_parcontract RS rtrancl_mono, | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 166 | parcontract_subset_reduce RS rtrancl_subset_rtrancl] 1)); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 167 | qed "reduce_eq_parreduce"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 168 | |
| 5069 | 169 | Goal "diamond(contract^*)"; | 
| 4089 | 170 | by (simp_tac (simpset() addsimps [reduce_eq_parreduce, diamond_rtrancl, | 
| 3120 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 171 | diamond_parcontract]) 1); | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 172 | qed "diamond_reduce"; | 
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 173 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 174 | |
| 
c58423c20740
New directory to contain examples of (co)inductive definitions
 paulson parents: diff
changeset | 175 | writeln"Reached end of file."; |