(* Title: HOL/Extraction/Higman.thy ID: $Id: Higman.thy,v 1.26 2007/11/13 09:53:39 berghofe Exp $ Author: Stefan Berghofer, TU Muenchen Monika Seisenberger, LMU Muenchen *) header {* Higman's lemma *} theory Higman imports Main begin text {* Formalization by Stefan Berghofer and Monika Seisenberger, based on Coquand and Fridlender \cite{Coquand93}. *} datatype letter = A | B inductive emb :: "letter list => letter list => bool" where emb0 [Pure.intro]: "emb [] bs" | emb1 [Pure.intro]: "emb as bs ==> emb as (b # bs)" | emb2 [Pure.intro]: "emb as bs ==> emb (a # as) (a # bs)" inductive L :: "letter list => letter list list => bool" for v :: "letter list" where L0 [Pure.intro]: "emb w v ==> L v (w # ws)" | L1 [Pure.intro]: "L v ws ==> L v (w # ws)" inductive good :: "letter list list => bool" where good0 [Pure.intro]: "L w ws ==> good (w # ws)" | good1 [Pure.intro]: "good ws ==> good (w # ws)" inductive R :: "letter => letter list list => letter list list => bool" for a :: letter where R0 [Pure.intro]: "R a [] []" | R1 [Pure.intro]: "R a vs ws ==> R a (w # vs) ((a # w) # ws)" inductive T :: "letter => letter list list => letter list list => bool" for a :: letter where T0 [Pure.intro]: "a ≠ b ==> R b ws zs ==> T a (w # zs) ((a # w) # zs)" | T1 [Pure.intro]: "T a ws zs ==> T a (w # ws) ((a # w) # zs)" | T2 [Pure.intro]: "a ≠ b ==> T a ws zs ==> T a ws ((b # w) # zs)" inductive bar :: "letter list list => bool" where bar1 [Pure.intro]: "good ws ==> bar ws" | bar2 [Pure.intro]: "(!!w. bar (w # ws)) ==> bar ws" theorem prop1: "bar ([] # ws)" by iprover theorem lemma1: "L as ws ==> L (a # as) ws" by (erule L.induct, iprover+) lemma lemma2': "R a vs ws ==> L as vs ==> L (a # as) ws" apply (induct set: R) apply (erule L.cases) apply simp+ apply (erule L.cases) apply simp_all apply (rule L0) apply (erule emb2) apply (erule L1) done lemma lemma2: "R a vs ws ==> good vs ==> good ws" apply (induct set: R) apply iprover apply (erule good.cases) apply simp_all apply (rule good0) apply (erule lemma2') apply assumption apply (erule good1) done lemma lemma3': "T a vs ws ==> L as vs ==> L (a # as) ws" apply (induct set: T) apply (erule L.cases) apply simp_all apply (rule L0) apply (erule emb2) apply (rule L1) apply (erule lemma1) apply (erule L.cases) apply simp_all apply iprover+ done lemma lemma3: "T a ws zs ==> good ws ==> good zs" apply (induct set: T) apply (erule good.cases) apply simp_all apply (rule good0) apply (erule lemma1) apply (erule good1) apply (erule good.cases) apply simp_all apply (rule good0) apply (erule lemma3') apply iprover+ done lemma lemma4: "R a ws zs ==> ws ≠ [] ==> T a ws zs" apply (induct set: R) apply iprover apply (case_tac vs) apply (erule R.cases) apply simp apply (case_tac a) apply (rule_tac b=B in T0) apply simp apply (rule R0) apply (rule_tac b=A in T0) apply simp apply (rule R0) apply simp apply (rule T1) apply simp done lemma letter_neq: "(a::letter) ≠ b ==> c ≠ a ==> c = b" apply (case_tac a) apply (case_tac b) apply (case_tac c, simp, simp) apply (case_tac c, simp, simp) apply (case_tac b) apply (case_tac c, simp, simp) apply (case_tac c, simp, simp) done lemma letter_eq_dec: "(a::letter) = b ∨ a ≠ b" apply (case_tac a) apply (case_tac b) apply simp apply simp apply (case_tac b) apply simp apply simp done theorem prop2: assumes ab: "a ≠ b" and bar: "bar xs" shows "!!ys zs. bar ys ==> T a xs zs ==> T b ys zs ==> bar zs" using bar proof induct fix xs zs assume "T a xs zs" and "good xs" hence "good zs" by (rule lemma3) then show "bar zs" by (rule bar1) next fix xs ys assume I: "!!w ys zs. bar ys ==> T a (w # xs) zs ==> T b ys zs ==> bar zs" assume "bar ys" thus "!!zs. T a xs zs ==> T b ys zs ==> bar zs" proof induct fix ys zs assume "T b ys zs" and "good ys" then have "good zs" by (rule lemma3) then show "bar zs" by (rule bar1) next fix ys zs assume I': "!!w zs. T a xs zs ==> T b (w # ys) zs ==> bar zs" and ys: "!!w. bar (w # ys)" and Ta: "T a xs zs" and Tb: "T b ys zs" show "bar zs" proof (rule bar2) fix w show "bar (w # zs)" proof (cases w) case Nil thus ?thesis by simp (rule prop1) next case (Cons c cs) from letter_eq_dec show ?thesis proof assume ca: "c = a" from ab have "bar ((a # cs) # zs)" by (iprover intro: I ys Ta Tb) thus ?thesis by (simp add: Cons ca) next assume "c ≠ a" with ab have cb: "c = b" by (rule letter_neq) from ab have "bar ((b # cs) # zs)" by (iprover intro: I' Ta Tb) thus ?thesis by (simp add: Cons cb) qed qed qed qed qed theorem prop3: assumes bar: "bar xs" shows "!!zs. xs ≠ [] ==> R a xs zs ==> bar zs" using bar proof induct fix xs zs assume "R a xs zs" and "good xs" then have "good zs" by (rule lemma2) then show "bar zs" by (rule bar1) next fix xs zs assume I: "!!w zs. w # xs ≠ [] ==> R a (w # xs) zs ==> bar zs" and xsb: "!!w. bar (w # xs)" and xsn: "xs ≠ []" and R: "R a xs zs" show "bar zs" proof (rule bar2) fix w show "bar (w # zs)" proof (induct w) case Nil show ?case by (rule prop1) next case (Cons c cs) from letter_eq_dec show ?case proof assume "c = a" thus ?thesis by (iprover intro: I [simplified] R) next from R xsn have T: "T a xs zs" by (rule lemma4) assume "c ≠ a" thus ?thesis by (iprover intro: prop2 Cons xsb xsn R T) qed qed qed qed theorem higman: "bar []" proof (rule bar2) fix w show "bar [w]" proof (induct w) show "bar [[]]" by (rule prop1) next fix c cs assume "bar [cs]" thus "bar [c # cs]" by (rule prop3) (simp, iprover) qed qed consts is_prefix :: "'a list => (nat => 'a) => bool" primrec "is_prefix [] f = True" "is_prefix (x # xs) f = (x = f (length xs) ∧ is_prefix xs f)" theorem L_idx: assumes L: "L w ws" shows "is_prefix ws f ==> ∃i. emb (f i) w ∧ i < length ws" using L proof induct case (L0 v ws) hence "emb (f (length ws)) w" by simp moreover have "length ws < length (v # ws)" by simp ultimately show ?case by iprover next case (L1 ws v) then obtain i where emb: "emb (f i) w" and "i < length ws" by simp iprover hence "i < length (v # ws)" by simp with emb show ?case by iprover qed theorem good_idx: assumes good: "good ws" shows "is_prefix ws f ==> ∃i j. emb (f i) (f j) ∧ i < j" using good proof induct case (good0 w ws) hence "w = f (length ws)" and "is_prefix ws f" by simp_all with good0 show ?case by (iprover dest: L_idx) next case (good1 ws w) thus ?case by simp qed theorem bar_idx: assumes bar: "bar ws" shows "is_prefix ws f ==> ∃i j. emb (f i) (f j) ∧ i < j" using bar proof induct case (bar1 ws) thus ?case by (rule good_idx) next case (bar2 ws) hence "is_prefix (f (length ws) # ws) f" by simp thus ?case by (rule bar2) qed text {* Strong version: yields indices of words that can be embedded into each other. *} theorem higman_idx: "∃(i::nat) j. emb (f i) (f j) ∧ i < j" proof (rule bar_idx) show "bar []" by (rule higman) show "is_prefix [] f" by simp qed text {* Weak version: only yield sequence containing words that can be embedded into each other. *} theorem good_prefix_lemma: assumes bar: "bar ws" shows "is_prefix ws f ==> ∃vs. is_prefix vs f ∧ good vs" using bar proof induct case bar1 thus ?case by iprover next case (bar2 ws) from bar2.prems have "is_prefix (f (length ws) # ws) f" by simp thus ?case by (iprover intro: bar2) qed theorem good_prefix: "∃vs. is_prefix vs f ∧ good vs" using higman by (rule good_prefix_lemma) simp+ subsection {* Extracting the program *} declare R.induct [ind_realizer] declare T.induct [ind_realizer] declare L.induct [ind_realizer] declare good.induct [ind_realizer] declare bar.induct [ind_realizer] extract higman_idx text {* Program extracted from the proof of @{text higman_idx}: @{thm [display] higman_idx_def [no_vars]} Corresponding correctness theorem: @{thm [display] higman_idx_correctness [no_vars]} Program extracted from the proof of @{text higman}: @{thm [display] higman_def [no_vars]} Program extracted from the proof of @{text prop1}: @{thm [display] prop1_def [no_vars]} Program extracted from the proof of @{text prop2}: @{thm [display] prop2_def [no_vars]} Program extracted from the proof of @{text prop3}: @{thm [display] prop3_def [no_vars]} *} subsection {* Some examples *} (* an attempt to express examples in HOL -- function mk_word :: "nat => randseed => letter list × randseed" where "mk_word k = (do i \<leftarrow> random 10; (if i > 7 ∧ k > 2 ∨ k > 1000 then return [] else do let l = (if i mod 2 = 0 then A else B); ls \<leftarrow> mk_word (Suc k); return (l # ls) done) done)" by pat_completeness auto termination by (relation "measure ((op -) 1001)") auto fun mk_word' :: "nat => randseed => letter list × randseed" where "mk_word' 0 = mk_word 0" | "mk_word' (Suc n) = (do _ \<leftarrow> mk_word 0; mk_word' n done)"*) consts_code "arbitrary :: LT" ("({* L0 [] [] *})") "arbitrary :: TT" ("({* T0 A [] [] [] R0 *})") code_module Higman contains higman = higman_idx ML {* local open Higman in val a = 16807.0; val m = 2147483647.0; fun nextRand seed = let val t = a*seed in t - m * real (Real.floor(t/m)) end; fun mk_word seed l = let val r = nextRand seed; val i = Real.round (r / m * 10.0); in if i > 7 andalso l > 2 then (r, []) else apsnd (cons (if i mod 2 = 0 then A else B)) (mk_word r (l+1)) end; fun f s zero = mk_word s 0 | f s (Suc n) = f (fst (mk_word s 0)) n; val g1 = snd o (f 20000.0); val g2 = snd o (f 50000.0); fun f1 zero = [A,A] | f1 (Suc zero) = [B] | f1 (Suc (Suc zero)) = [A,B] | f1 _ = []; fun f2 zero = [A,A] | f2 (Suc zero) = [B] | f2 (Suc (Suc zero)) = [B,A] | f2 _ = []; val (i1, j1) = higman g1; val (v1, w1) = (g1 i1, g1 j1); val (i2, j2) = higman g2; val (v2, w2) = (g2 i2, g2 j2); val (i3, j3) = higman f1; val (v3, w3) = (f1 i3, f1 j3); val (i4, j4) = higman f2; val (v4, w4) = (f2 i4, f2 j4); end; *} definition arbitrary_LT :: LT where [symmetric, code inline]: "arbitrary_LT = arbitrary" definition arbitrary_TT :: TT where [symmetric, code inline]: "arbitrary_TT = arbitrary" code_datatype L0 L1 arbitrary_LT code_datatype T0 T1 T2 arbitrary_TT export_code higman_idx in SML module_name Higman ML {* local open Higman in val a = 16807.0; val m = 2147483647.0; fun nextRand seed = let val t = a*seed in t - m * real (Real.floor(t/m)) end; fun mk_word seed l = let val r = nextRand seed; val i = Real.round (r / m * 10.0); in if i > 7 andalso l > 2 then (r, []) else apsnd (cons (if i mod 2 = 0 then A else B)) (mk_word r (l+1)) end; fun f s Zero_nat = mk_word s 0 | f s (Suc n) = f (fst (mk_word s 0)) n; val g1 = snd o (f 20000.0); val g2 = snd o (f 50000.0); fun f1 Zero_nat = [A,A] | f1 (Suc Zero_nat) = [B] | f1 (Suc (Suc Zero_nat)) = [A,B] | f1 _ = []; fun f2 Zero_nat = [A,A] | f2 (Suc Zero_nat) = [B] | f2 (Suc (Suc Zero_nat)) = [B,A] | f2 _ = []; val (i1, j1) = higman_idx g1; val (v1, w1) = (g1 i1, g1 j1); val (i2, j2) = higman_idx g2; val (v2, w2) = (g2 i2, g2 j2); val (i3, j3) = higman_idx f1; val (v3, w3) = (f1 i3, f1 j3); val (i4, j4) = higman_idx f2; val (v4, w4) = (f2 i4, f2 j4); end; *} end
theorem prop1:
bar ([] # ws)
theorem lemma1:
L as ws ==> L (a # as) ws
lemma lemma2':
[| R a vs ws; L as vs |] ==> L (a # as) ws
lemma lemma2:
[| R a vs ws; good vs |] ==> good ws
lemma lemma3':
[| T a vs ws; L as vs |] ==> L (a # as) ws
lemma lemma3:
[| T a ws zs; good ws |] ==> good zs
lemma lemma4:
[| R a ws zs; ws ≠ [] |] ==> T a ws zs
lemma letter_neq:
[| a ≠ b; c ≠ a |] ==> c = b
lemma letter_eq_dec:
a = b ∨ a ≠ b
theorem prop2:
[| a ≠ b; bar xs; bar ys; T a xs zs; T b ys zs |] ==> bar zs
theorem prop3:
[| bar xs; xs ≠ []; R a xs zs |] ==> bar zs
theorem higman:
bar []
theorem L_idx:
[| L w ws; is_prefix ws f |] ==> ∃i. emb (f i) w ∧ i < length ws
theorem good_idx:
[| good ws; is_prefix ws f |] ==> ∃i j. emb (f i) (f j) ∧ i < j
theorem bar_idx:
[| bar ws; is_prefix ws f |] ==> ∃i j. emb (f i) (f j) ∧ i < j
theorem higman_idx:
∃i j. emb (f i) (f j) ∧ i < j
theorem good_prefix_lemma:
[| bar ws; is_prefix ws f |] ==> ∃vs. is_prefix vs f ∧ good vs
theorem good_prefix:
∃vs. is_prefix vs f ∧ good vs