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