src/ZF/ex/Comb.ML
changeset 3014 f5554654d211
parent 2954 a16e1011bcc1
child 4091 771b1f6422a8
equal deleted inserted replaced
3013:01a563785367 3014:f5554654d211
    14 HOL system proofs may be found in
    14 HOL system proofs may be found in
    15 /usr/groups/theory/hvg-aftp/contrib/rule-induction/cl.ml
    15 /usr/groups/theory/hvg-aftp/contrib/rule-induction/cl.ml
    16 *)
    16 *)
    17 
    17 
    18 open Comb;
    18 open Comb;
    19 
       
    20 val [_,_,comb_Ap] = comb.intrs;
       
    21 val Ap_E = comb.mk_cases comb.con_defs "p#q : comb";
       
    22 
       
    23 
       
    24 (*** Results about Contraction ***)
       
    25 
       
    26 (*For type checking: replaces a-1->b by a,b:comb *)
       
    27 val contract_combE2 = contract.dom_subset RS subsetD RS SigmaE2;
       
    28 val contract_combD1 = contract.dom_subset RS subsetD RS SigmaD1;
       
    29 val contract_combD2 = contract.dom_subset RS subsetD RS SigmaD2;
       
    30 
       
    31 goal Comb.thy "field(contract) = comb";
       
    32 by (blast_tac (!claset addIs [contract.K] addSEs [contract_combE2]) 1);
       
    33 qed "field_contract_eq";
       
    34 
       
    35 bind_thm ("reduction_refl",
       
    36     (field_contract_eq RS equalityD2 RS subsetD RS rtrancl_refl));
       
    37 
       
    38 bind_thm ("rtrancl_into_rtrancl2",
       
    39     (r_into_rtrancl RS (trans_rtrancl RS transD)));
       
    40 
       
    41 val reduction_rls = [reduction_refl, contract.K, contract.S, 
       
    42                      contract.K RS rtrancl_into_rtrancl2,
       
    43                      contract.S RS rtrancl_into_rtrancl2,
       
    44                      contract.Ap1 RS rtrancl_into_rtrancl2,
       
    45                      contract.Ap2 RS rtrancl_into_rtrancl2];
       
    46 
       
    47 (*Example only: not used*)
       
    48 goalw Comb.thy [I_def] "!!p. p:comb ==> I#p ---> p";
       
    49 by (REPEAT (ares_tac (comb.intrs @ reduction_rls) 1));
       
    50 qed "reduce_I";
       
    51 
       
    52 goalw Comb.thy [I_def] "I: comb";
       
    53 by (REPEAT (ares_tac comb.intrs 1));
       
    54 qed "comb_I";
       
    55 
       
    56 (** Non-contraction results **)
       
    57 
       
    58 (*Derive a case for each combinator constructor*)
       
    59 val K_contractE = contract.mk_cases comb.con_defs "K -1-> r";
       
    60 val S_contractE = contract.mk_cases comb.con_defs "S -1-> r";
       
    61 val Ap_contractE = contract.mk_cases comb.con_defs "p#q -1-> r";
       
    62 
       
    63 AddSIs comb.intrs;
       
    64 AddIs  contract.intrs;
       
    65 AddSEs [contract_combD1,contract_combD2];     (*type checking*)
       
    66 AddSEs [K_contractE, S_contractE, Ap_contractE];
       
    67 AddSEs comb.free_SEs;
       
    68 
       
    69 goalw Comb.thy [I_def] "!!r. I -1-> r ==> P";
       
    70 by (Blast_tac 1);
       
    71 qed "I_contract_E";
       
    72 
       
    73 goal Comb.thy "!!p r. K#p -1-> r ==> (EX q. r = K#q & p -1-> q)";
       
    74 by (Blast_tac 1);
       
    75 qed "K1_contractD";
       
    76 
       
    77 goal Comb.thy "!!p r. [| p ---> q;  r: comb |] ==> p#r ---> q#r";
       
    78 by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
       
    79 by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
       
    80 by (etac rtrancl_induct 1);
       
    81 by (blast_tac (!claset addIs reduction_rls) 1);
       
    82 by (etac (trans_rtrancl RS transD) 1);
       
    83 by (fast_tac (!claset addIs reduction_rls) 1);
       
    84 qed "Ap_reduce1";
       
    85 
       
    86 goal Comb.thy "!!p r. [| p ---> q;  r: comb |] ==> r#p ---> r#q";
       
    87 by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
       
    88 by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
       
    89 by (etac rtrancl_induct 1);
       
    90 by (blast_tac (!claset addIs reduction_rls) 1);
       
    91 by (etac (trans_rtrancl RS transD) 1);
       
    92 by (fast_tac (!claset addIs reduction_rls) 1);
       
    93 qed "Ap_reduce2";
       
    94 
       
    95 (** Counterexample to the diamond property for -1-> **)
       
    96 
       
    97 goal Comb.thy "K#I#(I#I) -1-> I";
       
    98 by (REPEAT (ares_tac [contract.K, comb_I, comb_Ap] 1));
       
    99 qed "KIII_contract1";
       
   100 
       
   101 goalw Comb.thy [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))";
       
   102 by (DEPTH_SOLVE (resolve_tac (comb.intrs @ contract.intrs) 1));
       
   103 qed "KIII_contract2";
       
   104 
       
   105 goal Comb.thy "K#I#((K#I)#(K#I)) -1-> I";
       
   106 by (REPEAT (ares_tac (comb.intrs @ [contract.K, comb_I]) 1));
       
   107 qed "KIII_contract3";
       
   108 
       
   109 goalw Comb.thy [diamond_def] "~ diamond(contract)";
       
   110 by (blast_tac (!claset addIs [KIII_contract1,KIII_contract2,KIII_contract3]
       
   111                     addSEs [I_contract_E]) 1);
       
   112 qed "not_diamond_contract";
       
   113 
       
   114 
       
   115 
       
   116 (*** Results about Parallel Contraction ***)
       
   117 
       
   118 (*For type checking: replaces a=1=>b by a,b:comb *)
       
   119 val parcontract_combE2 = parcontract.dom_subset RS subsetD RS SigmaE2;
       
   120 val parcontract_combD1 = parcontract.dom_subset RS subsetD RS SigmaD1;
       
   121 val parcontract_combD2 = parcontract.dom_subset RS subsetD RS SigmaD2;
       
   122 
       
   123 goal Comb.thy "field(parcontract) = comb";
       
   124 by (blast_tac (!claset addIs [parcontract.K] 
       
   125                       addSEs [parcontract_combE2]) 1);
       
   126 qed "field_parcontract_eq";
       
   127 
       
   128 (*Derive a case for each combinator constructor*)
       
   129 val K_parcontractE = parcontract.mk_cases comb.con_defs "K =1=> r";
       
   130 val S_parcontractE = parcontract.mk_cases comb.con_defs "S =1=> r";
       
   131 val Ap_parcontractE = parcontract.mk_cases comb.con_defs "p#q =1=> r";
       
   132 
       
   133 AddSIs comb.intrs;
       
   134 AddIs  parcontract.intrs;
       
   135 AddSEs [Ap_E, K_parcontractE, S_parcontractE, Ap_parcontractE];
       
   136 AddSEs [parcontract_combD1, parcontract_combD2];     (*type checking*)
       
   137 AddSEs comb.free_SEs;
       
   138 
       
   139 (*** Basic properties of parallel contraction ***)
       
   140 
       
   141 goal Comb.thy "!!p r. K#p =1=> r ==> (EX p'. r = K#p' & p =1=> p')";
       
   142 by (Blast_tac 1);
       
   143 qed "K1_parcontractD";
       
   144 
       
   145 goal Comb.thy "!!p r. S#p =1=> r ==> (EX p'. r = S#p' & p =1=> p')";
       
   146 by (Blast_tac 1);
       
   147 qed "S1_parcontractD";
       
   148 
       
   149 goal Comb.thy
       
   150  "!!p q r. S#p#q =1=> r ==> (EX p' q'. r = S#p'#q' & p =1=> p' & q =1=> q')";
       
   151 by (blast_tac (!claset addSDs [S1_parcontractD]) 1);
       
   152 qed "S2_parcontractD";
       
   153 
       
   154 (*Church-Rosser property for parallel contraction*)
       
   155 goalw Comb.thy [diamond_def] "diamond(parcontract)";
       
   156 by (rtac (impI RS allI RS allI) 1);
       
   157 by (etac parcontract.induct 1);
       
   158 by (ALLGOALS 
       
   159     (fast_tac (!claset addSDs [K1_parcontractD,S2_parcontractD])));
       
   160 qed "diamond_parcontract";
       
   161 
    19 
   162 (*** Transitive closure preserves the Church-Rosser property ***)
    20 (*** Transitive closure preserves the Church-Rosser property ***)
   163 
    21 
   164 goalw Comb.thy [diamond_def]
    22 goalw Comb.thy [diamond_def]
   165     "!!x y r. [| diamond(r);  <x,y>:r^+ |] ==> \
    23     "!!x y r. [| diamond(r);  <x,y>:r^+ |] ==> \
   172 
    30 
   173 val [major] = goal Comb.thy "diamond(r) ==> diamond(r^+)";
    31 val [major] = goal Comb.thy "diamond(r) ==> diamond(r^+)";
   174 by (rewtac diamond_def);  (*unfold only in goal, not in premise!*)
    32 by (rewtac diamond_def);  (*unfold only in goal, not in premise!*)
   175 by (rtac (impI RS allI RS allI) 1);
    33 by (rtac (impI RS allI RS allI) 1);
   176 by (etac trancl_induct 1);
    34 by (etac trancl_induct 1);
   177 by (ALLGOALS
    35 by (ALLGOALS (*Seems to be a brittle, undirected search*)
   178     (slow_best_tac (!claset addIs [r_into_trancl, trans_trancl RS transD]
    36     (slow_best_tac (!claset addIs [r_into_trancl, trans_trancl RS transD]
   179                           addEs [major RS diamond_strip_lemmaE])));
    37                             addEs [major RS diamond_strip_lemmaE])));
   180 qed "diamond_trancl";
    38 qed "diamond_trancl";
   181 
    39 
       
    40 
       
    41 val [_,_,comb_Ap] = comb.intrs;
       
    42 val Ap_E = comb.mk_cases comb.con_defs "p#q : comb";
       
    43 
       
    44 AddSIs comb.intrs;
       
    45 AddSEs comb.free_SEs;
       
    46 
       
    47 
       
    48 (*** Results about Contraction ***)
       
    49 
       
    50 (*For type checking: replaces a-1->b by a,b:comb *)
       
    51 val contract_combE2 = contract.dom_subset RS subsetD RS SigmaE2;
       
    52 val contract_combD1 = contract.dom_subset RS subsetD RS SigmaD1;
       
    53 val contract_combD2 = contract.dom_subset RS subsetD RS SigmaD2;
       
    54 
       
    55 goal Comb.thy "field(contract) = comb";
       
    56 by (blast_tac (!claset addIs [contract.K] addSEs [contract_combE2]) 1);
       
    57 qed "field_contract_eq";
       
    58 
       
    59 bind_thm ("reduction_refl",
       
    60     (field_contract_eq RS equalityD2 RS subsetD RS rtrancl_refl));
       
    61 
       
    62 bind_thm ("rtrancl_into_rtrancl2",
       
    63     (r_into_rtrancl RS (trans_rtrancl RS transD)));
       
    64 
       
    65 AddSIs [reduction_refl, contract.K, contract.S];
       
    66 
       
    67 val reduction_rls = [contract.K RS rtrancl_into_rtrancl2,
       
    68                      contract.S RS rtrancl_into_rtrancl2,
       
    69                      contract.Ap1 RS rtrancl_into_rtrancl2,
       
    70                      contract.Ap2 RS rtrancl_into_rtrancl2];
       
    71 
       
    72 (*Example only: not used*)
       
    73 goalw Comb.thy [I_def] "!!p. p:comb ==> I#p ---> p";
       
    74 by (blast_tac (!claset addIs reduction_rls) 1);
       
    75 qed "reduce_I";
       
    76 
       
    77 goalw Comb.thy [I_def] "I: comb";
       
    78 by (Blast_tac 1);
       
    79 qed "comb_I";
       
    80 
       
    81 (** Non-contraction results **)
       
    82 
       
    83 (*Derive a case for each combinator constructor*)
       
    84 val K_contractE = contract.mk_cases comb.con_defs "K -1-> r";
       
    85 val S_contractE = contract.mk_cases comb.con_defs "S -1-> r";
       
    86 val Ap_contractE = contract.mk_cases comb.con_defs "p#q -1-> r";
       
    87 
       
    88 AddIs  [contract.Ap1, contract.Ap2];
       
    89 AddSEs [K_contractE, S_contractE, Ap_contractE];
       
    90 
       
    91 goalw Comb.thy [I_def] "!!r. I -1-> r ==> P";
       
    92 by (Blast_tac 1);
       
    93 qed "I_contract_E";
       
    94 
       
    95 goal Comb.thy "!!p r. K#p -1-> r ==> (EX q. r = K#q & p -1-> q)";
       
    96 by (Blast_tac 1);
       
    97 qed "K1_contractD";
       
    98 
       
    99 goal Comb.thy "!!p r. [| p ---> q;  r: comb |] ==> p#r ---> q#r";
       
   100 by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
       
   101 by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
       
   102 by (etac rtrancl_induct 1);
       
   103 by (blast_tac (!claset addIs reduction_rls) 1);
       
   104 by (etac (trans_rtrancl RS transD) 1);
       
   105 by (blast_tac (!claset addIs (contract_combD2::reduction_rls)) 1);
       
   106 qed "Ap_reduce1";
       
   107 
       
   108 goal Comb.thy "!!p r. [| p ---> q;  r: comb |] ==> r#p ---> r#q";
       
   109 by (forward_tac [rtrancl_type RS subsetD RS SigmaD1] 1);
       
   110 by (dtac (field_contract_eq RS equalityD1 RS subsetD) 1);
       
   111 by (etac rtrancl_induct 1);
       
   112 by (blast_tac (!claset addIs reduction_rls) 1);
       
   113 by (etac (trans_rtrancl RS transD) 1);
       
   114 by (blast_tac (!claset addIs (contract_combD2::reduction_rls)) 1);
       
   115 qed "Ap_reduce2";
       
   116 
       
   117 (** Counterexample to the diamond property for -1-> **)
       
   118 
       
   119 goal Comb.thy "K#I#(I#I) -1-> I";
       
   120 by (REPEAT (ares_tac [contract.K, comb_I, comb_Ap] 1));
       
   121 qed "KIII_contract1";
       
   122 
       
   123 goalw Comb.thy [I_def] "K#I#(I#I) -1-> K#I#((K#I)#(K#I))";
       
   124 by (DEPTH_SOLVE (resolve_tac (comb.intrs @ contract.intrs) 1));
       
   125 qed "KIII_contract2";
       
   126 
       
   127 goal Comb.thy "K#I#((K#I)#(K#I)) -1-> I";
       
   128 by (REPEAT (ares_tac (comb.intrs @ [contract.K, comb_I]) 1));
       
   129 qed "KIII_contract3";
       
   130 
       
   131 goalw Comb.thy [diamond_def] "~ diamond(contract)";
       
   132 by (blast_tac (!claset addIs [KIII_contract1,KIII_contract2,KIII_contract3]
       
   133                        addSEs [I_contract_E]) 1);
       
   134 qed "not_diamond_contract";
       
   135 
       
   136 
       
   137 
       
   138 (*** Results about Parallel Contraction ***)
       
   139 
       
   140 (*For type checking: replaces a=1=>b by a,b:comb *)
       
   141 val parcontract_combE2 = parcontract.dom_subset RS subsetD RS SigmaE2;
       
   142 val parcontract_combD1 = parcontract.dom_subset RS subsetD RS SigmaD1;
       
   143 val parcontract_combD2 = parcontract.dom_subset RS subsetD RS SigmaD2;
       
   144 
       
   145 goal Comb.thy "field(parcontract) = comb";
       
   146 by (blast_tac (!claset addIs [parcontract.K] 
       
   147                       addSEs [parcontract_combE2]) 1);
       
   148 qed "field_parcontract_eq";
       
   149 
       
   150 (*Derive a case for each combinator constructor*)
       
   151 val K_parcontractE = parcontract.mk_cases comb.con_defs "K =1=> r";
       
   152 val S_parcontractE = parcontract.mk_cases comb.con_defs "S =1=> r";
       
   153 val Ap_parcontractE = parcontract.mk_cases comb.con_defs "p#q =1=> r";
       
   154 
       
   155 AddIs  parcontract.intrs;
       
   156 AddSEs [Ap_E, K_parcontractE, S_parcontractE, Ap_parcontractE];
       
   157 
       
   158 (*** Basic properties of parallel contraction ***)
       
   159 
       
   160 goal Comb.thy "!!p r. K#p =1=> r ==> (EX p'. r = K#p' & p =1=> p')";
       
   161 by (Blast_tac 1);
       
   162 qed "K1_parcontractD";
       
   163 
       
   164 goal Comb.thy "!!p r. S#p =1=> r ==> (EX p'. r = S#p' & p =1=> p')";
       
   165 by (Blast_tac 1);
       
   166 qed "S1_parcontractD";
       
   167 
       
   168 goal Comb.thy
       
   169  "!!p q r. S#p#q =1=> r ==> (EX p' q'. r = S#p'#q' & p =1=> p' & q =1=> q')";
       
   170 by (blast_tac (!claset addSDs [S1_parcontractD]) 1);
       
   171 qed "S2_parcontractD";
       
   172 
       
   173 (*Church-Rosser property for parallel contraction*)
       
   174 goalw Comb.thy [diamond_def] "diamond(parcontract)";
       
   175 by (rtac (impI RS allI RS allI) 1);
       
   176 by (etac parcontract.induct 1);
       
   177 by (ALLGOALS 
       
   178     (blast_tac (!claset addSDs [K1_parcontractD, S2_parcontractD]
       
   179                         addIs  [parcontract_combD2])));
       
   180 qed "diamond_parcontract";
   182 
   181 
   183 (*** Equivalence of p--->q and p===>q ***)
   182 (*** Equivalence of p--->q and p===>q ***)
   184 
   183 
   185 goal Comb.thy "!!p q. p-1->q ==> p=1=>q";
   184 goal Comb.thy "!!p q. p-1->q ==> p=1=>q";
   186 by (etac contract.induct 1);
   185 by (etac contract.induct 1);
   200 goal Comb.thy "!!p q. p=1=>q ==> p--->q";
   199 goal Comb.thy "!!p q. p=1=>q ==> p--->q";
   201 by (etac parcontract.induct 1);
   200 by (etac parcontract.induct 1);
   202 by (blast_tac (!claset addIs reduction_rls) 1);
   201 by (blast_tac (!claset addIs reduction_rls) 1);
   203 by (blast_tac (!claset addIs reduction_rls) 1);
   202 by (blast_tac (!claset addIs reduction_rls) 1);
   204 by (blast_tac (!claset addIs reduction_rls) 1);
   203 by (blast_tac (!claset addIs reduction_rls) 1);
   205 by (rtac (trans_rtrancl RS transD) 1);
   204 by (blast_tac 
   206 by (ALLGOALS 
   205     (!claset addIs [trans_rtrancl RS transD,
   207     (blast_tac 
   206 		    Ap_reduce1, Ap_reduce2, parcontract_combD1,
   208      (!claset addIs [Ap_reduce1, Ap_reduce2, parcontract_combD1,
   207 		    parcontract_combD2]) 1);
   209 		     parcontract_combD2])));
       
   210 qed "parcontract_imp_reduce";
   208 qed "parcontract_imp_reduce";
   211 
   209 
   212 goal Comb.thy "!!p q. p===>q ==> p--->q";
   210 goal Comb.thy "!!p q. p===>q ==> p--->q";
   213 by (forward_tac [trancl_type RS subsetD RS SigmaD1] 1);
   211 by (forward_tac [trancl_type RS subsetD RS SigmaD1] 1);
   214 by (dtac (field_parcontract_eq RS equalityD1 RS subsetD) 1);
   212 by (dtac (field_parcontract_eq RS equalityD1 RS subsetD) 1);