| 12560 |      1 | (*  Title:      ZF/Induct/Comb.thy
 | 
| 12088 |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Lawrence C Paulson
 | 
|  |      4 |     Copyright   1994  University of Cambridge
 | 
|  |      5 | *)
 | 
|  |      6 | 
 | 
| 12610 |      7 | header {* Combinatory Logic example: the Church-Rosser Theorem *}
 | 
| 12088 |      8 | 
 | 
| 16417 |      9 | theory Comb imports Main begin
 | 
| 12088 |     10 | 
 | 
| 12610 |     11 | text {*
 | 
|  |     12 |   Curiously, combinators do not include free variables.
 | 
|  |     13 | 
 | 
|  |     14 |   Example taken from \cite{camilleri-melham}.
 | 
|  |     15 | *}
 | 
|  |     16 | 
 | 
|  |     17 | subsection {* Definitions *}
 | 
|  |     18 | 
 | 
|  |     19 | text {* Datatype definition of combinators @{text S} and @{text K}. *}
 | 
| 12088 |     20 | 
 | 
| 12610 |     21 | consts comb :: i
 | 
|  |     22 | datatype comb =
 | 
|  |     23 |     K
 | 
|  |     24 |   | S
 | 
| 15775 |     25 |   | app ("p \<in> comb", "q \<in> comb")    (infixl "@@" 90)
 | 
| 12610 |     26 | 
 | 
|  |     27 | text {*
 | 
|  |     28 |   Inductive definition of contractions, @{text "-1->"} and
 | 
|  |     29 |   (multi-step) reductions, @{text "--->"}.
 | 
|  |     30 | *}
 | 
|  |     31 | 
 | 
| 12088 |     32 | consts
 | 
| 12610 |     33 |   contract  :: i
 | 
|  |     34 | syntax
 | 
| 15775 |     35 |   "_contract"       :: "[i,i] => o"    (infixl "-1->" 50)
 | 
| 12610 |     36 |   "_contract_multi" :: "[i,i] => o"    (infixl "--->" 50)
 | 
| 12088 |     37 | translations
 | 
| 12560 |     38 |   "p -1-> q" == "<p,q> \<in> contract"
 | 
|  |     39 |   "p ---> q" == "<p,q> \<in> contract^*"
 | 
| 12088 |     40 | 
 | 
| 15775 |     41 | syntax (xsymbols)
 | 
| 15863 |     42 |   "comb.app"    :: "[i, i] => i"    	     (infixl "\<bullet>" 90)
 | 
| 15775 |     43 | 
 | 
| 12088 |     44 | inductive
 | 
| 12610 |     45 |   domains "contract" \<subseteq> "comb \<times> comb"
 | 
| 12560 |     46 |   intros
 | 
| 15775 |     47 |     K:   "[| p \<in> comb;  q \<in> comb |] ==> K\<bullet>p\<bullet>q -1-> p"
 | 
|  |     48 |     S:   "[| p \<in> comb;  q \<in> comb;  r \<in> comb |] ==> S\<bullet>p\<bullet>q\<bullet>r -1-> (p\<bullet>r)\<bullet>(q\<bullet>r)"
 | 
|  |     49 |     Ap1: "[| p-1->q;  r \<in> comb |] ==> p\<bullet>r -1-> q\<bullet>r"
 | 
|  |     50 |     Ap2: "[| p-1->q;  r \<in> comb |] ==> r\<bullet>p -1-> r\<bullet>q"
 | 
| 12610 |     51 |   type_intros comb.intros
 | 
| 12088 |     52 | 
 | 
| 12610 |     53 | text {*
 | 
|  |     54 |   Inductive definition of parallel contractions, @{text "=1=>"} and
 | 
|  |     55 |   (multi-step) parallel reductions, @{text "===>"}.
 | 
|  |     56 | *}
 | 
|  |     57 | 
 | 
| 12088 |     58 | consts
 | 
| 12610 |     59 |   parcontract :: i
 | 
|  |     60 | syntax
 | 
|  |     61 |   "_parcontract" :: "[i,i] => o"    (infixl "=1=>" 50)
 | 
|  |     62 |   "_parcontract_multi" :: "[i,i] => o"    (infixl "===>" 50)
 | 
| 12088 |     63 | translations
 | 
| 12560 |     64 |   "p =1=> q" == "<p,q> \<in> parcontract"
 | 
|  |     65 |   "p ===> q" == "<p,q> \<in> parcontract^+"
 | 
| 12088 |     66 | 
 | 
|  |     67 | inductive
 | 
| 12610 |     68 |   domains "parcontract" \<subseteq> "comb \<times> comb"
 | 
| 12560 |     69 |   intros
 | 
| 12610 |     70 |     refl: "[| p \<in> comb |] ==> p =1=> p"
 | 
| 15775 |     71 |     K:    "[| p \<in> comb;  q \<in> comb |] ==> K\<bullet>p\<bullet>q =1=> p"
 | 
|  |     72 |     S:    "[| p \<in> comb;  q \<in> comb;  r \<in> comb |] ==> S\<bullet>p\<bullet>q\<bullet>r =1=> (p\<bullet>r)\<bullet>(q\<bullet>r)"
 | 
|  |     73 |     Ap:   "[| p=1=>q;  r=1=>s |] ==> p\<bullet>r =1=> q\<bullet>s"
 | 
| 12610 |     74 |   type_intros comb.intros
 | 
| 12088 |     75 | 
 | 
| 12610 |     76 | text {*
 | 
|  |     77 |   Misc definitions.
 | 
|  |     78 | *}
 | 
| 12088 |     79 | 
 | 
| 24894 |     80 | definition
 | 
|  |     81 |   I :: i  where
 | 
| 15775 |     82 |   "I == S\<bullet>K\<bullet>K"
 | 
| 12088 |     83 | 
 | 
| 24894 |     84 | definition
 | 
|  |     85 |   diamond :: "i => o"  where
 | 
| 12610 |     86 |   "diamond(r) ==
 | 
|  |     87 |     \<forall>x y. <x,y>\<in>r --> (\<forall>y'. <x,y'>\<in>r --> (\<exists>z. <y,z>\<in>r & <y',z> \<in> r))"
 | 
| 12560 |     88 | 
 | 
|  |     89 | 
 | 
| 12610 |     90 | subsection {* Transitive closure preserves the Church-Rosser property *}
 | 
| 12560 |     91 | 
 | 
|  |     92 | lemma diamond_strip_lemmaD [rule_format]:
 | 
| 12610 |     93 |   "[| diamond(r);  <x,y>:r^+ |] ==>
 | 
|  |     94 |     \<forall>y'. <x,y'>:r --> (\<exists>z. <y',z>: r^+ & <y,z>: r)"
 | 
|  |     95 |   apply (unfold diamond_def)
 | 
|  |     96 |   apply (erule trancl_induct)
 | 
|  |     97 |    apply (blast intro: r_into_trancl)
 | 
|  |     98 |   apply clarify
 | 
|  |     99 |   apply (drule spec [THEN mp], assumption)
 | 
|  |    100 |   apply (blast intro: r_into_trancl trans_trancl [THEN transD])
 | 
|  |    101 |   done
 | 
| 12560 |    102 | 
 | 
|  |    103 | lemma diamond_trancl: "diamond(r) ==> diamond(r^+)"
 | 
| 12610 |    104 |   apply (simp (no_asm_simp) add: diamond_def)
 | 
|  |    105 |   apply (rule impI [THEN allI, THEN allI])
 | 
|  |    106 |   apply (erule trancl_induct)
 | 
|  |    107 |    apply auto
 | 
|  |    108 |    apply (best intro: r_into_trancl trans_trancl [THEN transD]
 | 
|  |    109 |      dest: diamond_strip_lemmaD)+
 | 
|  |    110 |   done
 | 
| 12560 |    111 | 
 | 
| 15775 |    112 | inductive_cases Ap_E [elim!]: "p\<bullet>q \<in> comb"
 | 
| 12560 |    113 | 
 | 
|  |    114 | declare comb.intros [intro!]
 | 
|  |    115 | 
 | 
|  |    116 | 
 | 
| 12610 |    117 | subsection {* Results about Contraction *}
 | 
| 12560 |    118 | 
 | 
| 12610 |    119 | text {*
 | 
|  |    120 |   For type checking: replaces @{term "a -1-> b"} by @{text "a, b \<in>
 | 
|  |    121 |   comb"}.
 | 
|  |    122 | *}
 | 
|  |    123 | 
 | 
| 12560 |    124 | lemmas contract_combE2 = contract.dom_subset [THEN subsetD, THEN SigmaE2]
 | 
| 12610 |    125 |   and contract_combD1 = contract.dom_subset [THEN subsetD, THEN SigmaD1]
 | 
|  |    126 |   and contract_combD2 = contract.dom_subset [THEN subsetD, THEN SigmaD2]
 | 
| 12560 |    127 | 
 | 
|  |    128 | lemma field_contract_eq: "field(contract) = comb"
 | 
| 12610 |    129 |   by (blast intro: contract.K elim!: contract_combE2)
 | 
| 12560 |    130 | 
 | 
| 12610 |    131 | lemmas reduction_refl =
 | 
|  |    132 |   field_contract_eq [THEN equalityD2, THEN subsetD, THEN rtrancl_refl]
 | 
| 12560 |    133 | 
 | 
| 12610 |    134 | lemmas rtrancl_into_rtrancl2 =
 | 
|  |    135 |   r_into_rtrancl [THEN trans_rtrancl [THEN transD]]
 | 
| 12560 |    136 | 
 | 
|  |    137 | declare reduction_refl [intro!] contract.K [intro!] contract.S [intro!]
 | 
|  |    138 | 
 | 
| 12610 |    139 | lemmas reduction_rls =
 | 
|  |    140 |   contract.K [THEN rtrancl_into_rtrancl2]
 | 
|  |    141 |   contract.S [THEN rtrancl_into_rtrancl2]
 | 
|  |    142 |   contract.Ap1 [THEN rtrancl_into_rtrancl2]
 | 
|  |    143 |   contract.Ap2 [THEN rtrancl_into_rtrancl2]
 | 
| 12560 |    144 | 
 | 
| 15775 |    145 | lemma "p \<in> comb ==> I\<bullet>p ---> p"
 | 
| 12610 |    146 |   -- {* Example only: not used *}
 | 
|  |    147 |   by (unfold I_def) (blast intro: reduction_rls)
 | 
| 12560 |    148 | 
 | 
|  |    149 | lemma comb_I: "I \<in> comb"
 | 
| 12610 |    150 |   by (unfold I_def) blast
 | 
| 12560 |    151 | 
 | 
| 12610 |    152 | 
 | 
|  |    153 | subsection {* Non-contraction results *}
 | 
| 12560 |    154 | 
 | 
| 12610 |    155 | text {* Derive a case for each combinator constructor. *}
 | 
|  |    156 | 
 | 
|  |    157 | inductive_cases
 | 
|  |    158 |       K_contractE [elim!]: "K -1-> r"
 | 
|  |    159 |   and S_contractE [elim!]: "S -1-> r"
 | 
| 15775 |    160 |   and Ap_contractE [elim!]: "p\<bullet>q -1-> r"
 | 
| 12560 |    161 | 
 | 
|  |    162 | lemma I_contract_E: "I -1-> r ==> P"
 | 
| 12610 |    163 |   by (auto simp add: I_def)
 | 
| 12560 |    164 | 
 | 
| 15775 |    165 | lemma K1_contractD: "K\<bullet>p -1-> r ==> (\<exists>q. r = K\<bullet>q & p -1-> q)"
 | 
| 12610 |    166 |   by auto
 | 
| 12560 |    167 | 
 | 
| 15775 |    168 | lemma Ap_reduce1: "[| p ---> q;  r \<in> comb |] ==> p\<bullet>r ---> q\<bullet>r"
 | 
| 12610 |    169 |   apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1])
 | 
|  |    170 |   apply (drule field_contract_eq [THEN equalityD1, THEN subsetD])
 | 
|  |    171 |   apply (erule rtrancl_induct)
 | 
|  |    172 |    apply (blast intro: reduction_rls)
 | 
|  |    173 |   apply (erule trans_rtrancl [THEN transD])
 | 
|  |    174 |   apply (blast intro: contract_combD2 reduction_rls)
 | 
|  |    175 |   done
 | 
| 12560 |    176 | 
 | 
| 15775 |    177 | lemma Ap_reduce2: "[| p ---> q;  r \<in> comb |] ==> r\<bullet>p ---> r\<bullet>q"
 | 
| 12610 |    178 |   apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1])
 | 
|  |    179 |   apply (drule field_contract_eq [THEN equalityD1, THEN subsetD])
 | 
|  |    180 |   apply (erule rtrancl_induct)
 | 
|  |    181 |    apply (blast intro: reduction_rls)
 | 
| 13573 |    182 |   apply (blast intro: trans_rtrancl [THEN transD] 
 | 
|  |    183 |                       contract_combD2 reduction_rls)
 | 
| 12610 |    184 |   done
 | 
| 12560 |    185 | 
 | 
| 12610 |    186 | text {* Counterexample to the diamond property for @{text "-1->"}. *}
 | 
| 12560 |    187 | 
 | 
| 15775 |    188 | lemma KIII_contract1: "K\<bullet>I\<bullet>(I\<bullet>I) -1-> I"
 | 
| 12610 |    189 |   by (blast intro: comb.intros contract.K comb_I)
 | 
| 12560 |    190 | 
 | 
| 15775 |    191 | lemma KIII_contract2: "K\<bullet>I\<bullet>(I\<bullet>I) -1-> K\<bullet>I\<bullet>((K\<bullet>I)\<bullet>(K\<bullet>I))"
 | 
| 12610 |    192 |   by (unfold I_def) (blast intro: comb.intros contract.intros)
 | 
| 12560 |    193 | 
 | 
| 15775 |    194 | lemma KIII_contract3: "K\<bullet>I\<bullet>((K\<bullet>I)\<bullet>(K\<bullet>I)) -1-> I"
 | 
| 12610 |    195 |   by (blast intro: comb.intros contract.K comb_I)
 | 
| 12560 |    196 | 
 | 
| 12610 |    197 | lemma not_diamond_contract: "\<not> diamond(contract)"
 | 
|  |    198 |   apply (unfold diamond_def)
 | 
|  |    199 |   apply (blast intro: KIII_contract1 KIII_contract2 KIII_contract3
 | 
|  |    200 |     elim!: I_contract_E)
 | 
|  |    201 |   done
 | 
| 12560 |    202 | 
 | 
|  |    203 | 
 | 
| 12610 |    204 | subsection {* Results about Parallel Contraction *}
 | 
| 12560 |    205 | 
 | 
| 12610 |    206 | text {* For type checking: replaces @{text "a =1=> b"} by @{text "a, b
 | 
|  |    207 |   \<in> comb"} *}
 | 
| 12560 |    208 | lemmas parcontract_combE2 = parcontract.dom_subset [THEN subsetD, THEN SigmaE2]
 | 
| 12610 |    209 |   and parcontract_combD1 = parcontract.dom_subset [THEN subsetD, THEN SigmaD1]
 | 
|  |    210 |   and parcontract_combD2 = parcontract.dom_subset [THEN subsetD, THEN SigmaD2]
 | 
| 12560 |    211 | 
 | 
|  |    212 | lemma field_parcontract_eq: "field(parcontract) = comb"
 | 
| 12610 |    213 |   by (blast intro: parcontract.K elim!: parcontract_combE2)
 | 
| 12560 |    214 | 
 | 
| 12610 |    215 | text {* Derive a case for each combinator constructor. *}
 | 
|  |    216 | inductive_cases
 | 
|  |    217 |       K_parcontractE [elim!]: "K =1=> r"
 | 
|  |    218 |   and S_parcontractE [elim!]: "S =1=> r"
 | 
| 15775 |    219 |   and Ap_parcontractE [elim!]: "p\<bullet>q =1=> r"
 | 
| 12560 |    220 | 
 | 
|  |    221 | declare parcontract.intros [intro]
 | 
|  |    222 | 
 | 
|  |    223 | 
 | 
| 12610 |    224 | subsection {* Basic properties of parallel contraction *}
 | 
| 12560 |    225 | 
 | 
| 12610 |    226 | lemma K1_parcontractD [dest!]:
 | 
| 15775 |    227 |     "K\<bullet>p =1=> r ==> (\<exists>p'. r = K\<bullet>p' & p =1=> p')"
 | 
| 12610 |    228 |   by auto
 | 
| 12560 |    229 | 
 | 
| 12610 |    230 | lemma S1_parcontractD [dest!]:
 | 
| 15775 |    231 |     "S\<bullet>p =1=> r ==> (\<exists>p'. r = S\<bullet>p' & p =1=> p')"
 | 
| 12610 |    232 |   by auto
 | 
| 12560 |    233 | 
 | 
|  |    234 | lemma S2_parcontractD [dest!]:
 | 
| 15775 |    235 |     "S\<bullet>p\<bullet>q =1=> r ==> (\<exists>p' q'. r = S\<bullet>p'\<bullet>q' & p =1=> p' & q =1=> q')"
 | 
| 12610 |    236 |   by auto
 | 
| 12560 |    237 | 
 | 
|  |    238 | lemma diamond_parcontract: "diamond(parcontract)"
 | 
| 12610 |    239 |   -- {* Church-Rosser property for parallel contraction *}
 | 
|  |    240 |   apply (unfold diamond_def)
 | 
|  |    241 |   apply (rule impI [THEN allI, THEN allI])
 | 
|  |    242 |   apply (erule parcontract.induct)
 | 
|  |    243 |      apply (blast elim!: comb.free_elims  intro: parcontract_combD2)+
 | 
|  |    244 |   done
 | 
| 12560 |    245 | 
 | 
| 12610 |    246 | text {*
 | 
|  |    247 |   \medskip Equivalence of @{prop "p ---> q"} and @{prop "p ===> q"}.
 | 
|  |    248 | *}
 | 
| 12560 |    249 | 
 | 
|  |    250 | lemma contract_imp_parcontract: "p-1->q ==> p=1=>q"
 | 
| 18415 |    251 |   by (induct set: contract) auto
 | 
| 12560 |    252 | 
 | 
|  |    253 | lemma reduce_imp_parreduce: "p--->q ==> p===>q"
 | 
| 12610 |    254 |   apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1])
 | 
|  |    255 |   apply (drule field_contract_eq [THEN equalityD1, THEN subsetD])
 | 
|  |    256 |   apply (erule rtrancl_induct)
 | 
|  |    257 |    apply (blast intro: r_into_trancl)
 | 
|  |    258 |   apply (blast intro: contract_imp_parcontract r_into_trancl
 | 
|  |    259 |     trans_trancl [THEN transD])
 | 
|  |    260 |   done
 | 
| 12560 |    261 | 
 | 
|  |    262 | lemma parcontract_imp_reduce: "p=1=>q ==> p--->q"
 | 
| 18415 |    263 |   apply (induct set: parcontract)
 | 
| 12610 |    264 |      apply (blast intro: reduction_rls)
 | 
|  |    265 |     apply (blast intro: reduction_rls)
 | 
| 12560 |    266 |    apply (blast intro: reduction_rls)
 | 
| 12610 |    267 |   apply (blast intro: trans_rtrancl [THEN transD]
 | 
|  |    268 |     Ap_reduce1 Ap_reduce2 parcontract_combD1 parcontract_combD2)
 | 
|  |    269 |   done
 | 
| 12560 |    270 | 
 | 
|  |    271 | lemma parreduce_imp_reduce: "p===>q ==> p--->q"
 | 
| 12610 |    272 |   apply (frule trancl_type [THEN subsetD, THEN SigmaD1])
 | 
|  |    273 |   apply (drule field_parcontract_eq [THEN equalityD1, THEN subsetD])
 | 
|  |    274 |   apply (erule trancl_induct, erule parcontract_imp_reduce)
 | 
|  |    275 |   apply (erule trans_rtrancl [THEN transD])
 | 
|  |    276 |   apply (erule parcontract_imp_reduce)
 | 
|  |    277 |   done
 | 
| 12560 |    278 | 
 | 
|  |    279 | lemma parreduce_iff_reduce: "p===>q <-> p--->q"
 | 
| 12610 |    280 |   by (blast intro: parreduce_imp_reduce reduce_imp_parreduce)
 | 
| 12088 |    281 | 
 | 
|  |    282 | end
 |