diff --git a/src/HOL/Induct/Comb.thy b/src/HOL/Induct/Comb.thy --- a/src/HOL/Induct/Comb.thy +++ b/src/HOL/Induct/Comb.thy @@ -1,196 +1,201 @@ (* Title: HOL/Induct/Comb.thy Author: Lawrence C Paulson Copyright 1996 University of Cambridge *) section \Combinatory Logic example: the Church-Rosser Theorem\ theory Comb imports Main begin text \ - Curiously, combinators do not include free variables. - + Combinator terms do not have free variables. Example taken from @{cite camilleri92}. \ - subsection \Definitions\ text \Datatype definition of combinators \S\ and \K\.\ datatype comb = K | S | Ap comb comb (infixl "\" 90) text \ Inductive definition of contractions, \\\<^sup>1\ and (multi-step) reductions, \\\. \ -inductive_set contract :: "(comb*comb) set" - and contract_rel1 :: "[comb,comb] \ bool" (infixl "\\<^sup>1" 50) -where - "x \\<^sup>1 y == (x,y) \ contract" -| K: "K\x\y \\<^sup>1 x" -| S: "S\x\y\z \\<^sup>1 (x\z)\(y\z)" -| Ap1: "x\\<^sup>1y \ x\z \\<^sup>1 y\z" -| Ap2: "x\\<^sup>1y \ z\x \\<^sup>1 z\y" +inductive contract1 :: "[comb,comb] \ bool" (infixl "\\<^sup>1" 50) + where + K: "K\x\y \\<^sup>1 x" + | S: "S\x\y\z \\<^sup>1 (x\z)\(y\z)" + | Ap1: "x \\<^sup>1 y \ x\z \\<^sup>1 y\z" + | Ap2: "x \\<^sup>1 y \ z\x \\<^sup>1 z\y" abbreviation - contract_rel :: "[comb,comb] \ bool" (infixl "\" 50) where - "x \ y == (x,y) \ contract\<^sup>*" + contract :: "[comb,comb] \ bool" (infixl "\" 50) where + "contract \ contract1\<^sup>*\<^sup>*" text \ Inductive definition of parallel contractions, \\\<^sup>1\ and (multi-step) parallel reductions, \\\. \ -inductive_set parcontract :: "(comb*comb) set" - and parcontract_rel1 :: "[comb,comb] \ bool" (infixl "\\<^sup>1" 50) -where - "x \\<^sup>1 y == (x,y) \ parcontract" -| refl: "x \\<^sup>1 x" -| K: "K\x\y \\<^sup>1 x" -| S: "S\x\y\z \\<^sup>1 (x\z)\(y\z)" -| Ap: "[| x\\<^sup>1y; z\\<^sup>1w |] ==> x\z \\<^sup>1 y\w" +inductive parcontract1 :: "[comb,comb] \ bool" (infixl "\\<^sup>1" 50) + where + refl: "x \\<^sup>1 x" + | K: "K\x\y \\<^sup>1 x" + | S: "S\x\y\z \\<^sup>1 (x\z)\(y\z)" + | Ap: "\x \\<^sup>1 y; z \\<^sup>1 w\ \ x\z \\<^sup>1 y\w" abbreviation - parcontract_rel :: "[comb,comb] \ bool" (infixl "\" 50) where - "x \ y == (x,y) \ parcontract\<^sup>*" + parcontract :: "[comb,comb] \ bool" (infixl "\" 50) where + "parcontract \ parcontract1\<^sup>*\<^sup>*" text \ Misc definitions. \ definition I :: comb where - "I = S\K\K" + "I \ S\K\K" definition - diamond :: "('a * 'a)set \ bool" where + diamond :: "([comb,comb] \ bool) \ bool" where \ \confluence; Lambda/Commutation treats this more abstractly\ - "diamond(r) = (\x y. (x,y) \ r --> - (\y'. (x,y') \ r --> - (\z. (y,z) \ r & (y',z) \ r)))" + "diamond r \ \x y. r x y \ + (\y'. r x y' \ + (\z. r y z \ r y' z))" subsection \Reflexive/Transitive closure preserves Church-Rosser property\ -text\So does the Transitive closure, with a similar proof\ +text\Remark: So does the Transitive closure, with a similar proof\ text\Strip lemma. The induction hypothesis covers all but the last diamond of the strip.\ -lemma diamond_strip_lemmaE [rule_format]: - "[| diamond(r); (x,y) \ r\<^sup>* |] ==> - \y'. (x,y') \ r --> (\z. (y',z) \ r\<^sup>* & (y,z) \ r)" -apply (unfold diamond_def) -apply (erule rtrancl_induct) -apply (meson rtrancl_refl) -apply (meson rtrancl_trans r_into_rtrancl) -done +lemma strip_lemma [rule_format]: + assumes "diamond r" and r: "r\<^sup>*\<^sup>* x y" "r x y'" + shows "\z. r\<^sup>*\<^sup>* y' z \ r y z" + using r +proof (induction rule: rtranclp_induct) + case base + then show ?case + by blast +next + case (step y z) + then show ?case + using \diamond r\ unfolding diamond_def + by (metis rtranclp.rtrancl_into_rtrancl) +qed -lemma diamond_rtrancl: "diamond(r) \ diamond(r\<^sup>*)" -apply (simp (no_asm_simp) add: diamond_def) -apply (rule impI [THEN allI, THEN allI]) -apply (erule rtrancl_induct, blast) -apply (meson rtrancl_trans r_into_rtrancl diamond_strip_lemmaE) -done +proposition diamond_rtrancl: + assumes "diamond r" + shows "diamond(r\<^sup>*\<^sup>*)" + unfolding diamond_def +proof (intro strip) + fix x y y' + assume "r\<^sup>*\<^sup>* x y" "r\<^sup>*\<^sup>* x y'" + then show "\z. r\<^sup>*\<^sup>* y z \ r\<^sup>*\<^sup>* y' z" + proof (induction rule: rtranclp_induct) + case base + then show ?case + by blast + next + case (step y z) + then show ?case + by (meson assms strip_lemma rtranclp.rtrancl_into_rtrancl) + qed +qed subsection \Non-contraction results\ text \Derive a case for each combinator constructor.\ inductive_cases - K_contractE [elim!]: "K \\<^sup>1 r" + K_contractE [elim!]: "K \\<^sup>1 r" and S_contractE [elim!]: "S \\<^sup>1 r" and Ap_contractE [elim!]: "p\q \\<^sup>1 r" -declare contract.K [intro!] contract.S [intro!] -declare contract.Ap1 [intro] contract.Ap2 [intro] +declare contract1.K [intro!] contract1.S [intro!] +declare contract1.Ap1 [intro] contract1.Ap2 [intro] -lemma I_contract_E [elim!]: "I \\<^sup>1 z \ P" -by (unfold I_def, blast) +lemma I_contract_E [iff]: "\ I \\<^sup>1 z" + unfolding I_def by blast -lemma K1_contractD [elim!]: "K\x \\<^sup>1 z \ (\x'. z = K\x' & x \\<^sup>1 x')" -by blast +lemma K1_contractD [elim!]: "K\x \\<^sup>1 z \ (\x'. z = K\x' \ x \\<^sup>1 x')" + by blast lemma Ap_reduce1 [intro]: "x \ y \ x\z \ y\z" -apply (erule rtrancl_induct) -apply (blast intro: rtrancl_trans)+ -done + by (induction rule: rtranclp_induct; blast intro: rtranclp_trans) lemma Ap_reduce2 [intro]: "x \ y \ z\x \ z\y" -apply (erule rtrancl_induct) -apply (blast intro: rtrancl_trans)+ -done + by (induction rule: rtranclp_induct; blast intro: rtranclp_trans) text \Counterexample to the diamond property for \<^term>\x \\<^sup>1 y\\ -lemma not_diamond_contract: "~ diamond(contract)" -by (unfold diamond_def, metis S_contractE contract.K) +lemma not_diamond_contract: "\ diamond(contract1)" + unfolding diamond_def by (metis S_contractE contract1.K) subsection \Results about Parallel Contraction\ text \Derive a case for each combinator constructor.\ inductive_cases K_parcontractE [elim!]: "K \\<^sup>1 r" and S_parcontractE [elim!]: "S \\<^sup>1 r" and Ap_parcontractE [elim!]: "p\q \\<^sup>1 r" -declare parcontract.intros [intro] - -(*** Basic properties of parallel contraction ***) +declare parcontract1.intros [intro] subsection \Basic properties of parallel contraction\ - -lemma K1_parcontractD [dest!]: "K\x \\<^sup>1 z \ (\x'. z = K\x' & x \\<^sup>1 x')" -by blast +text\The rules below are not essential but make proofs much faster\ -lemma S1_parcontractD [dest!]: "S\x \\<^sup>1 z \ (\x'. z = S\x' & x \\<^sup>1 x')" -by blast +lemma K1_parcontractD [dest!]: "K\x \\<^sup>1 z \ (\x'. z = K\x' \ x \\<^sup>1 x')" + by blast -lemma S2_parcontractD [dest!]: - "S\x\y \\<^sup>1 z \ (\x' y'. z = S\x'\y' & x \\<^sup>1 x' & y \\<^sup>1 y')" -by blast +lemma S1_parcontractD [dest!]: "S\x \\<^sup>1 z \ (\x'. z = S\x' \ x \\<^sup>1 x')" + by blast -text\The rules above are not essential but make proofs much faster\ +lemma S2_parcontractD [dest!]: "S\x\y \\<^sup>1 z \ (\x' y'. z = S\x'\y' \ x \\<^sup>1 x' \ y \\<^sup>1 y')" + by blast text\Church-Rosser property for parallel contraction\ -lemma diamond_parcontract: "diamond parcontract" -apply (unfold diamond_def) -apply (rule impI [THEN allI, THEN allI]) -apply (erule parcontract.induct, fast+) -done +proposition diamond_parcontract: "diamond parcontract1" +proof - + have "(\z. w \\<^sup>1 z \ y' \\<^sup>1 z)" if "y \\<^sup>1 w" "y \\<^sup>1 y'" for w y y' + using that by (induction arbitrary: y' rule: parcontract1.induct) fast+ + then show ?thesis + by (auto simp: diamond_def) +qed -text \ - \<^medskip> - Equivalence of \<^prop>\p \ q\ and \<^prop>\p \ q\. -\ +subsection \Equivalence of \<^prop>\p \ q\ and \<^prop>\p \ q\.\ -lemma contract_subset_parcontract: "contract \ parcontract" -by (auto, erule contract.induct, blast+) +lemma contract_imp_parcontract: "x \\<^sup>1 y \ x \\<^sup>1 y" + by (induction rule: contract1.induct; blast) text\Reductions: simply throw together reflexivity, transitivity and the one-step reductions\ -declare r_into_rtrancl [intro] rtrancl_trans [intro] - -(*Example only: not used*) -lemma reduce_I: "I\x \ x" -by (unfold I_def, blast) +proposition reduce_I: "I\x \ x" + unfolding I_def + by (meson contract1.K contract1.S r_into_rtranclp rtranclp.rtrancl_into_rtrancl) -lemma parcontract_subset_reduce: "parcontract \ contract\<^sup>*" -by (auto, erule parcontract.induct, blast+) +lemma parcontract_imp_reduce: "x \\<^sup>1 y \ x \ y" +proof (induction rule: parcontract1.induct) + case (Ap x y z w) + then show ?case + by (meson Ap_reduce1 Ap_reduce2 rtranclp_trans) +qed auto -lemma reduce_eq_parreduce: "contract\<^sup>* = parcontract\<^sup>*" -by (metis contract_subset_parcontract parcontract_subset_reduce rtrancl_subset) +lemma reduce_eq_parreduce: "x \ y \ x \ y" + by (metis contract_imp_parcontract parcontract_imp_reduce predicate2I rtranclp_subset) -theorem diamond_reduce: "diamond(contract\<^sup>*)" -by (simp add: reduce_eq_parreduce diamond_rtrancl diamond_parcontract) +theorem diamond_reduce: "diamond(contract)" + using diamond_parcontract diamond_rtrancl reduce_eq_parreduce by presburger end