src/HOL/ex/Comb.ML
author berghofe
Tue, 25 Jun 1996 13:11:29 +0200
changeset 1824 44254696843a
parent 1820 e381e1c51689
child 1895 92b30c4829bf
permissions -rw-r--r--
Changed argument order of nat_rec.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    21
(*** Reflexive/Transitive closure preserves the Church-Rosser property 
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    22
     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
    23
***)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    24
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    25
val [_, spec_mp] = [spec] RL [mp];
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    26
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    27
(*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
    28
goalw Comb.thy [diamond_def]
1691
fbbd65c89c23 Fixed indenting
paulson
parents: 1689
diff changeset
    29
    "!!r. [| diamond(r);  (x,y):r^* |] ==> \ 
fbbd65c89c23 Fixed indenting
paulson
parents: 1689
diff changeset
    30
\         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
    31
by (etac rtrancl_induct 1);
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    32
by (fast_tac (!claset addIs [rtrancl_refl]) 1);
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    33
by (slow_best_tac (!claset addIs [r_into_rtrancl RSN (2, rtrancl_trans)]
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    34
                          addSDs [spec_mp]) 1);
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    35
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
    36
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    37
val [major] = goal Comb.thy "diamond(r) ==> diamond(r^*)";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    38
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
    39
by (rtac (impI RS allI RS allI) 1);
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    40
by (etac rtrancl_induct 1);
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    41
by (fast_tac (!claset addIs [rtrancl_refl]) 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    42
by (ALLGOALS
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    43
    (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
    44
                           addEs [major RS diamond_strip_lemmaE])));
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    45
qed "diamond_rtrancl";
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
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    48
(*** Results about Contraction ***)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    49
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    50
(** Non-contraction results **)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    51
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    52
(*Derive a case for each combinator constructor*)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    53
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
    54
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
    55
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
    56
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    57
AddIs  contract.intrs;
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    58
AddSEs [K_contractE, S_contractE, Ap_contractE];
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    59
Addss  (HOL_ss addsimps comb.simps);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    60
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    61
goalw Comb.thy [I_def] "!!z. I -1-> z ==> P";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    62
by (Fast_tac 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    63
qed "I_contract_E";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    64
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    65
(*Unused*)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    66
goal Comb.thy "!!x z. K#x -1-> z ==> (EX y. z = K#y & x -1-> y)";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    67
by (Fast_tac 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    68
qed "K1_contractD";
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);
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    72
by (rtac rtrancl_refl 1);
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    73
by (fast_tac (!claset addIs [r_into_rtrancl, rtrancl_trans]) 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    74
qed "Ap_reduce1";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    75
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    76
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
    77
by (etac rtrancl_induct 1);
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    78
by (rtac rtrancl_refl 1);
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    79
by (fast_tac (!claset addIs [r_into_rtrancl, rtrancl_trans]) 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    80
qed "Ap_reduce2";
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
(** Counterexample to the diamond property for -1-> **)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    83
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    84
goal Comb.thy "K#I#(I#I) -1-> I";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    85
by (rtac contract.K 1);
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    86
qed "KIII_contract1";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    87
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    88
goalw Comb.thy [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    89
by (Fast_tac 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    90
qed "KIII_contract2";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    91
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    92
goal Comb.thy "K#I#((K#I)#(K#I)) -1-> I";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    93
by (Fast_tac 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    94
qed "KIII_contract3";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    95
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    96
goalw Comb.thy [diamond_def] "~ diamond(contract)";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
    97
by (fast_tac (!claset addIs [KIII_contract1,KIII_contract2,KIII_contract3]
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    98
                     addSEs [I_contract_E]) 1);
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
    99
qed "not_diamond_contract";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   100
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
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   103
(*** Results about Parallel Contraction ***)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   104
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   105
(*Derive a case for each combinator constructor*)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   106
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
   107
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
   108
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
   109
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   110
AddIs  parcontract.intrs;
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   111
AddSEs [K_parcontractE, S_parcontractE,Ap_parcontractE];
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   112
Addss  (HOL_ss addsimps comb.simps);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   113
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   114
(*** Basic properties of parallel contraction ***)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   115
1689
c38d953caedb Fixed some unfortunate variable names
paulson
parents: 1639
diff changeset
   116
goal Comb.thy "!!x z. K#x =1=> z ==> (EX x'. z = K#x' & x =1=> x')";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   117
by (Fast_tac 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   118
qed "K1_parcontractD";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   119
1689
c38d953caedb Fixed some unfortunate variable names
paulson
parents: 1639
diff changeset
   120
goal Comb.thy "!!x z. S#x =1=> z ==> (EX x'. z = S#x' & x =1=> x')";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   121
by (Fast_tac 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   122
qed "S1_parcontractD";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   123
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   124
goal Comb.thy
1689
c38d953caedb Fixed some unfortunate variable names
paulson
parents: 1639
diff changeset
   125
 "!!x y z. S#x#y =1=> z ==> (EX x' y'. z = S#x'#y' & x =1=> x' & y =1=> y')";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   126
by (fast_tac (!claset addSDs [S1_parcontractD]) 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   127
     (*the extra rule makes the proof more than twice as fast*)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   128
qed "S2_parcontractD";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   129
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   130
(*Church-Rosser property for parallel contraction*)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   131
goalw Comb.thy [diamond_def] "diamond(parcontract)";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   132
by (rtac (impI RS allI RS allI) 1);
1739
35961ebbbfad Added prune_params_tac to improve readability of subgoals
paulson
parents: 1730
diff changeset
   133
by (etac parcontract.induct 1 THEN prune_params_tac);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   134
by (ALLGOALS 
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   135
    (fast_tac (!claset addSDs [K1_parcontractD,S2_parcontractD])));
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   136
qed "diamond_parcontract";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   137
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
(*** Equivalence of x--->y and x===>y ***)
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   140
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   141
goal Comb.thy "contract <= parcontract";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   142
by (rtac subsetI 1);
1730
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1691
diff changeset
   143
by (split_all_tac 1);
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   144
by (etac contract.induct 1);
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   145
by (ALLGOALS (Fast_tac));
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   146
qed "contract_subset_parcontract";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   147
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   148
(*Reductions: simply throw together reflexivity, transitivity and
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   149
  the one-step reductions*)
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   150
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   151
AddIs [rtrancl_refl, r_into_rtrancl,
1639
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   152
		       contract.K, contract.S, Ap_reduce1, Ap_reduce2,
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   153
		       rtrancl_trans];
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";
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   157
by (best_tac (!claset) 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
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1691
diff changeset
   162
by (split_all_tac 1);
1739
35961ebbbfad Added prune_params_tac to improve readability of subgoals
paulson
parents: 1730
diff changeset
   163
by (etac parcontract.induct 1 THEN prune_params_tac);
1820
e381e1c51689 Classical tactics now use default claset.
berghofe
parents: 1739
diff changeset
   164
by (ALLGOALS (deepen_tac (!claset) 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^*)";
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   175
by (simp_tac (HOL_ss addsimps [reduce_eq_parreduce, diamond_rtrancl, 
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   176
			       diamond_parcontract]) 1);
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   177
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   178
qed "diamond_reduce";
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
d3484e841d1e New example Comb: Church-Rosser for combinators, ported from ZF
paulson
parents:
diff changeset
   181
writeln"Reached end of file.";