| author | wenzelm | 
| Mon, 11 Sep 2000 17:59:14 +0200 | |
| changeset 9922 | ab4b408dbf96 | 
| parent 6144 | 7d38744313c8 | 
| child 11316 | b4e71bd751e4 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/ex/comb.ML | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 515 | 6 | Combinatory Logic example: the Church-Rosser Theorem | 
| 7 | Curiously, combinators do not include free variables. | |
| 8 | ||
| 9 | Example taken from | |
| 10 | J. Camilleri and T. F. Melham. | |
| 11 | Reasoning with Inductively Defined Relations in the HOL Theorem Prover. | |
| 12 | Report 265, University of Cambridge Computer Laboratory, 1992. | |
| 13 | ||
| 14 | HOL system proofs may be found in | |
| 15 | /usr/groups/theory/hvg-aftp/contrib/rule-induction/cl.ml | |
| 16 | *) | |
| 17 | ||
| 3014 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 18 | (*** Transitive closure preserves the Church-Rosser property ***) | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 19 | |
| 5068 | 20 | Goalw [diamond_def] | 
| 3014 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 21 | "!!x y r. [| diamond(r); <x,y>:r^+ |] ==> \ | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 22 | \ ALL y'. <x,y'>:r --> (EX z. <y',z>: r^+ & <y,z>: r)"; | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 23 | by (etac trancl_induct 1); | 
| 4091 | 24 | by (blast_tac (claset() addIs [r_into_trancl]) 1); | 
| 25 | by (slow_best_tac (claset() addSDs [spec RS mp] | |
| 3014 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 26 | addIs [r_into_trancl, trans_trancl RS transD]) 1); | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 27 | val diamond_strip_lemmaE = result() RS spec RS mp RS exE; | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 28 | |
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 29 | val [major] = goal Comb.thy "diamond(r) ==> diamond(r^+)"; | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 30 | by (rewtac diamond_def); (*unfold only in goal, not in premise!*) | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 31 | by (rtac (impI RS allI RS allI) 1); | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 32 | by (etac trancl_induct 1); | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 33 | by (ALLGOALS (*Seems to be a brittle, undirected search*) | 
| 4091 | 34 | (slow_best_tac (claset() addIs [r_into_trancl, trans_trancl RS transD] | 
| 3014 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 35 | addEs [major RS diamond_strip_lemmaE]))); | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 36 | qed "diamond_trancl"; | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 37 | |
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 38 | |
| 515 | 39 | val [_,_,comb_Ap] = comb.intrs; | 
| 6141 | 40 | val Ap_E = comb.mk_cases "p#q : comb"; | 
| 515 | 41 | |
| 3014 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 42 | AddSIs comb.intrs; | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 43 | |
| 515 | 44 | |
| 45 | (*** Results about Contraction ***) | |
| 46 | ||
| 47 | (*For type checking: replaces a-1->b by a,b:comb *) | |
| 48 | val contract_combE2 = contract.dom_subset RS subsetD RS SigmaE2; | |
| 49 | val contract_combD1 = contract.dom_subset RS subsetD RS SigmaD1; | |
| 50 | val contract_combD2 = contract.dom_subset RS subsetD RS SigmaD2; | |
| 51 | ||
| 5068 | 52 | Goal "field(contract) = comb"; | 
| 4091 | 53 | by (blast_tac (claset() addIs [contract.K] addSEs [contract_combE2]) 1); | 
| 760 | 54 | qed "field_contract_eq"; | 
| 515 | 55 | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 56 | bind_thm ("reduction_refl",
 | 
| 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 57 | (field_contract_eq RS equalityD2 RS subsetD RS rtrancl_refl)); | 
| 515 | 58 | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 59 | bind_thm ("rtrancl_into_rtrancl2",
 | 
| 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 60 | (r_into_rtrancl RS (trans_rtrancl RS transD))); | 
| 515 | 61 | |
| 3014 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 62 | AddSIs [reduction_refl, contract.K, contract.S]; | 
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 63 | |
| 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 64 | val reduction_rls = [contract.K RS rtrancl_into_rtrancl2, | 
| 1461 | 65 | contract.S RS rtrancl_into_rtrancl2, | 
| 66 | contract.Ap1 RS rtrancl_into_rtrancl2, | |
| 67 | contract.Ap2 RS rtrancl_into_rtrancl2]; | |
| 515 | 68 | |
| 69 | (*Example only: not used*) | |
| 5137 | 70 | Goalw [I_def] "p:comb ==> I#p ---> p"; | 
| 4091 | 71 | by (blast_tac (claset() addIs reduction_rls) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 72 | qed "reduce_I"; | 
| 515 | 73 | |
| 5068 | 74 | Goalw [I_def] "I: comb"; | 
| 3014 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 75 | by (Blast_tac 1); | 
| 760 | 76 | qed "comb_I"; | 
| 515 | 77 | |
| 78 | (** Non-contraction results **) | |
| 0 | 79 | |
| 515 | 80 | (*Derive a case for each combinator constructor*) | 
| 6141 | 81 | val K_contractE = contract.mk_cases "K -1-> r"; | 
| 82 | val S_contractE = contract.mk_cases "S -1-> r"; | |
| 83 | val Ap_contractE = contract.mk_cases "p#q -1-> r"; | |
| 515 | 84 | |
| 3014 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 85 | AddIs [contract.Ap1, contract.Ap2]; | 
| 2469 | 86 | AddSEs [K_contractE, S_contractE, Ap_contractE]; | 
| 515 | 87 | |
| 5137 | 88 | Goalw [I_def] "I -1-> r ==> P"; | 
| 6144 | 89 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 90 | qed "I_contract_E"; | 
| 515 | 91 | |
| 5137 | 92 | Goal "K#p -1-> r ==> (EX q. r = K#q & p -1-> q)"; | 
| 6144 | 93 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 94 | qed "K1_contractD"; | 
| 515 | 95 | |
| 5137 | 96 | Goal "[| p ---> q; r: comb |] ==> p#r ---> q#r"; | 
| 515 | 97 | by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1); | 
| 98 | by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1); | |
| 99 | by (etac rtrancl_induct 1); | |
| 4091 | 100 | by (blast_tac (claset() addIs reduction_rls) 1); | 
| 515 | 101 | by (etac (trans_rtrancl RS transD) 1); | 
| 4091 | 102 | by (blast_tac (claset() addIs (contract_combD2::reduction_rls)) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 103 | qed "Ap_reduce1"; | 
| 515 | 104 | |
| 5137 | 105 | Goal "[| p ---> q; r: comb |] ==> r#p ---> r#q"; | 
| 515 | 106 | by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1); | 
| 107 | by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1); | |
| 108 | by (etac rtrancl_induct 1); | |
| 4091 | 109 | by (blast_tac (claset() addIs reduction_rls) 1); | 
| 515 | 110 | by (etac (trans_rtrancl RS transD) 1); | 
| 4091 | 111 | by (blast_tac (claset() addIs (contract_combD2::reduction_rls)) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 112 | qed "Ap_reduce2"; | 
| 515 | 113 | |
| 114 | (** Counterexample to the diamond property for -1-> **) | |
| 115 | ||
| 5068 | 116 | Goal "K#I#(I#I) -1-> I"; | 
| 515 | 117 | by (REPEAT (ares_tac [contract.K, comb_I, comb_Ap] 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 118 | qed "KIII_contract1"; | 
| 515 | 119 | |
| 5068 | 120 | Goalw [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))"; | 
| 515 | 121 | by (DEPTH_SOLVE (resolve_tac (comb.intrs @ contract.intrs) 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 122 | qed "KIII_contract2"; | 
| 515 | 123 | |
| 5068 | 124 | Goal "K#I#((K#I)#(K#I)) -1-> I"; | 
| 515 | 125 | by (REPEAT (ares_tac (comb.intrs @ [contract.K, comb_I]) 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 126 | qed "KIII_contract3"; | 
| 515 | 127 | |
| 5068 | 128 | Goalw [diamond_def] "~ diamond(contract)"; | 
| 4091 | 129 | by (blast_tac (claset() addIs [KIII_contract1,KIII_contract2,KIII_contract3] | 
| 3014 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 130 | addSEs [I_contract_E]) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 131 | qed "not_diamond_contract"; | 
| 515 | 132 | |
| 0 | 133 | |
| 134 | ||
| 515 | 135 | (*** Results about Parallel Contraction ***) | 
| 136 | ||
| 137 | (*For type checking: replaces a=1=>b by a,b:comb *) | |
| 138 | val parcontract_combE2 = parcontract.dom_subset RS subsetD RS SigmaE2; | |
| 139 | val parcontract_combD1 = parcontract.dom_subset RS subsetD RS SigmaD1; | |
| 140 | val parcontract_combD2 = parcontract.dom_subset RS subsetD RS SigmaD2; | |
| 141 | ||
| 5068 | 142 | Goal "field(parcontract) = comb"; | 
| 4091 | 143 | by (blast_tac (claset() addIs [parcontract.K] | 
| 2496 | 144 | addSEs [parcontract_combE2]) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 145 | qed "field_parcontract_eq"; | 
| 515 | 146 | |
| 147 | (*Derive a case for each combinator constructor*) | |
| 6141 | 148 | val K_parcontractE = parcontract.mk_cases "K =1=> r"; | 
| 149 | val S_parcontractE = parcontract.mk_cases "S =1=> r"; | |
| 150 | val Ap_parcontractE = parcontract.mk_cases "p#q =1=> r"; | |
| 515 | 151 | |
| 2469 | 152 | AddIs parcontract.intrs; | 
| 153 | AddSEs [Ap_E, K_parcontractE, S_parcontractE, Ap_parcontractE]; | |
| 515 | 154 | |
| 155 | (*** Basic properties of parallel contraction ***) | |
| 156 | ||
| 5137 | 157 | Goal "K#p =1=> r ==> (EX p'. r = K#p' & p =1=> p')"; | 
| 6144 | 158 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 159 | qed "K1_parcontractD"; | 
| 6144 | 160 | AddSDs [K1_parcontractD]; | 
| 515 | 161 | |
| 5137 | 162 | Goal "S#p =1=> r ==> (EX p'. r = S#p' & p =1=> p')"; | 
| 6144 | 163 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 164 | qed "S1_parcontractD"; | 
| 6144 | 165 | AddSDs [S1_parcontractD]; | 
| 515 | 166 | |
| 5268 | 167 | Goal "S#p#q =1=> r ==> (EX p' q'. r = S#p'#q' & p =1=> p' & q =1=> q')"; | 
| 6144 | 168 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 169 | qed "S2_parcontractD"; | 
| 6144 | 170 | AddSDs [S2_parcontractD]; | 
| 515 | 171 | |
| 172 | (*Church-Rosser property for parallel contraction*) | |
| 5068 | 173 | Goalw [diamond_def] "diamond(parcontract)"; | 
| 515 | 174 | by (rtac (impI RS allI RS allI) 1); | 
| 1732 | 175 | by (etac parcontract.induct 1); | 
| 515 | 176 | by (ALLGOALS | 
| 6144 | 177 | (blast_tac (claset() addSEs comb.free_SEs | 
| 178 | addIs [parcontract_combD2]))); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 179 | qed "diamond_parcontract"; | 
| 515 | 180 | |
| 181 | (*** Equivalence of p--->q and p===>q ***) | |
| 182 | ||
| 5137 | 183 | Goal "p-1->q ==> p=1=>q"; | 
| 1732 | 184 | by (etac contract.induct 1); | 
| 6144 | 185 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 186 | qed "contract_imp_parcontract"; | 
| 0 | 187 | |
| 5137 | 188 | Goal "p--->q ==> p===>q"; | 
| 515 | 189 | by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1); | 
| 190 | by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1); | |
| 191 | by (etac rtrancl_induct 1); | |
| 4091 | 192 | by (blast_tac (claset() addIs [r_into_trancl]) 1); | 
| 193 | by (blast_tac (claset() addIs [contract_imp_parcontract, | |
| 6144 | 194 | r_into_trancl, trans_trancl RS transD]) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 195 | qed "reduce_imp_parreduce"; | 
| 515 | 196 | |
| 0 | 197 | |
| 5137 | 198 | Goal "p=1=>q ==> p--->q"; | 
| 1732 | 199 | by (etac parcontract.induct 1); | 
| 4091 | 200 | by (blast_tac (claset() addIs reduction_rls) 1); | 
| 201 | by (blast_tac (claset() addIs reduction_rls) 1); | |
| 202 | by (blast_tac (claset() addIs reduction_rls) 1); | |
| 3014 
f5554654d211
Moved diamond_trancl (which is independent of the rest) to the top
 paulson parents: 
2954diff
changeset | 203 | by (blast_tac | 
| 4091 | 204 | (claset() addIs [trans_rtrancl RS transD, | 
| 6144 | 205 | Ap_reduce1, Ap_reduce2, parcontract_combD1, | 
| 206 | parcontract_combD2]) 1); | |
| 760 | 207 | qed "parcontract_imp_reduce"; | 
| 515 | 208 | |
| 5137 | 209 | Goal "p===>q ==> p--->q"; | 
| 515 | 210 | by (forward_tac [trancl_type RS subsetD RS SigmaD1] 1); | 
| 211 | by (dtac (field_parcontract_eq RS equalityD1 RS subsetD) 1); | |
| 212 | by (etac trancl_induct 1); | |
| 213 | by (etac parcontract_imp_reduce 1); | |
| 214 | by (etac (trans_rtrancl RS transD) 1); | |
| 215 | by (etac parcontract_imp_reduce 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 216 | qed "parreduce_imp_reduce"; | 
| 515 | 217 | |
| 5068 | 218 | Goal "p===>q <-> p--->q"; | 
| 6144 | 219 | by (blast_tac (claset() addIs [parreduce_imp_reduce, reduce_imp_parreduce]) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 220 | qed "parreduce_iff_reduce"; | 
| 515 | 221 | |
| 222 | writeln"Reached end of file."; |