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