Up to index of Isabelle/HOL/SizeChange
theory Interpretation(* Title: HOL/Library/SCT_Interpretation.thy ID: $Id: Interpretation.thy,v 1.1 2007/11/06 16:44:54 krauss Exp $ Author: Alexander Krauss, TU Muenchen *) header {* Applying SCT to function definitions *} theory Interpretation imports Main Misc_Tools Criterion begin definition "idseq R s x = (s 0 = x ∧ (∀i. R (s (Suc i)) (s i)))" lemma not_acc_smaller: assumes notacc: "¬ accp R x" shows "∃y. R y x ∧ ¬ accp R y" proof (rule classical) assume "¬ ?thesis" hence "!!y. R y x ==> accp R y" by blast with accp.accI have "accp R x" . with notacc show ?thesis by contradiction qed lemma non_acc_has_idseq: assumes "¬ accp R x" shows "∃s. idseq R s x" proof - have "∃f. ∀x. ¬accp R x --> R (f x) x ∧ ¬accp R (f x)" by (rule choice, auto simp:not_acc_smaller) then obtain f where in_R: "!!x. ¬accp R x ==> R (f x) x" and nia: "!!x. ¬accp R x ==> ¬accp R (f x)" by blast let ?s = "λi. (f ^ i) x" { fix i have "¬accp R (?s i)" by (induct i) (auto simp:nia `¬accp R x`) hence "R (f (?s i)) (?s i)" by (rule in_R) } hence "idseq R ?s x" unfolding idseq_def by auto thus ?thesis by auto qed types ('a, 'q) cdesc = "('q => bool) × ('q => 'a) ×('q => 'a)" fun in_cdesc :: "('a, 'q) cdesc => 'a => 'a => bool" where "in_cdesc (Γ, r, l) x y = (∃q. x = r q ∧ y = l q ∧ Γ q)" fun mk_rel :: "('a, 'q) cdesc list => 'a => 'a => bool" where "mk_rel [] x y = False" | "mk_rel (c#cs) x y = (in_cdesc c x y ∨ mk_rel cs x y)" lemma some_rd: assumes "mk_rel rds x y" shows "∃rd∈set rds. in_cdesc rd x y" using assms by (induct rds) (auto simp:in_cdesc_def) (* from a value sequence, get a sequence of rds *) lemma ex_cs: assumes idseq: "idseq (mk_rel rds) s x" shows "∃cs. ∀i. cs i ∈ set rds ∧ in_cdesc (cs i) (s (Suc i)) (s i)" proof - from idseq have a: "∀i. ∃rd ∈ set rds. in_cdesc rd (s (Suc i)) (s i)" by (auto simp:idseq_def intro:some_rd) show ?thesis by (rule choice) (insert a, blast) qed types 'a measures = "nat => 'a => nat" fun stepP :: "('a, 'q) cdesc => ('a, 'q) cdesc => ('a => nat) => ('a => nat) => (nat => nat => bool) => bool" where "stepP (Γ1,r1,l1) (Γ2,r2,l2) m1 m2 R = (∀q1 q2. Γ1 q1 ∧ Γ2 q2 ∧ r1 q1 = l2 q2 --> R (m2 (l2 q2)) ((m1 (l1 q1))))" definition decr :: "('a, 'q) cdesc => ('a, 'q) cdesc => ('a => nat) => ('a => nat) => bool" where "decr c1 c2 m1 m2 = stepP c1 c2 m1 m2 (op <)" definition decreq :: "('a, 'q) cdesc => ('a, 'q) cdesc => ('a => nat) => ('a => nat) => bool" where "decreq c1 c2 m1 m2 = stepP c1 c2 m1 m2 (op ≤)" definition no_step :: "('a, 'q) cdesc => ('a, 'q) cdesc => bool" where "no_step c1 c2 = stepP c1 c2 (λx. 0) (λx. 0) (λx y. False)" lemma decr_in_cdesc: assumes "in_cdesc RD1 y x" assumes "in_cdesc RD2 z y" assumes "decr RD1 RD2 m1 m2" shows "m2 y < m1 x" using assms by (cases RD1, cases RD2, auto simp:decr_def) lemma decreq_in_cdesc: assumes "in_cdesc RD1 y x" assumes "in_cdesc RD2 z y" assumes "decreq RD1 RD2 m1 m2" shows "m2 y ≤ m1 x" using assms by (cases RD1, cases RD2, auto simp:decreq_def) lemma no_inf_desc_nat_sequence: fixes s :: "nat => nat" assumes leq: "!!i. n ≤ i ==> s (Suc i) ≤ s i" assumes less: "∃∞i. s (Suc i) < s i" shows False proof - { fix i j:: nat assume "n ≤ i" assume "i ≤ j" { fix k have "s (i + k) ≤ s i" proof (induct k) case 0 thus ?case by simp next case (Suc k) with leq[of "i + k"] `n ≤ i` show ?case by simp qed } from this[of "j - i"] `n ≤ i` `i ≤ j` have "s j ≤ s i" by auto } note decr = this let ?min = "LEAST x. x ∈ range (λi. s (n + i))" have "?min ∈ range (λi. s (n + i))" by (rule LeastI) auto then obtain k where min: "?min = s (n + k)" by auto from less obtain k' where "n + k < k'" and "s (Suc k') < s k'" unfolding INF_nat by auto with decr[of "n + k" k'] min have "s (Suc k') < ?min" by auto moreover from `n + k < k'` have "s (Suc k') = s (n + (Suc k' - n))" by simp ultimately show False using not_less_Least by blast qed definition approx :: "nat scg => ('a, 'q) cdesc => ('a, 'q) cdesc => 'a measures => 'a measures => bool" where "approx G C C' M M' = (∀i j. (dsc G i j --> decr C C' (M i) (M' j)) ∧(eq G i j --> decreq C C' (M i) (M' j)))" (* Unfolding "approx" for finite graphs *) lemma approx_empty: "approx (Graph {}) c1 c2 ms1 ms2" unfolding approx_def has_edge_def dest_graph.simps by simp lemma approx_less: assumes "stepP c1 c2 (ms1 i) (ms2 j) (op <)" assumes "approx (Graph Es) c1 c2 ms1 ms2" shows "approx (Graph (insert (i, \<down>, j) Es)) c1 c2 ms1 ms2" using assms unfolding approx_def has_edge_def dest_graph.simps decr_def by auto lemma approx_leq: assumes "stepP c1 c2 (ms1 i) (ms2 j) (op ≤)" assumes "approx (Graph Es) c1 c2 ms1 ms2" shows "approx (Graph (insert (i, \<Down>, j) Es)) c1 c2 ms1 ms2" using assms unfolding approx_def has_edge_def dest_graph.simps decreq_def by auto lemma "approx (Graph {(1, \<down>, 2),(2, \<Down>, 3)}) c1 c2 ms1 ms2" apply (intro approx_less approx_leq approx_empty) oops (* fun no_step :: "('a, 'q) cdesc => ('a, 'q) cdesc => bool" where "no_step (Γ1, r1, l1) (Γ2, r2, l2) = (∀q1 q2. Γ1 q1 ∧ Γ2 q2 ∧ r1 q1 = l2 q2 --> False)" *) lemma no_stepI: "stepP c1 c2 m1 m2 (λx y. False) ==> no_step c1 c2" by (cases c1, cases c2) (auto simp: no_step_def) definition sound_int :: "nat acg => ('a, 'q) cdesc list => 'a measures list => bool" where "sound_int \<A> RDs M = (∀n<length RDs. ∀m<length RDs. no_step (RDs ! n) (RDs ! m) ∨ (∃G. (\<A> \<turnstile> n \<leadsto>G m) ∧ approx G (RDs ! n) (RDs ! m) (M ! n) (M ! m)))" (* The following are uses by the tactics *) lemma length_simps: "length [] = 0" "length (x#xs) = Suc (length xs)" by auto lemma all_less_zero: "∀n<(0::nat). P n" by simp lemma all_less_Suc: assumes Pk: "P k" assumes Pn: "∀n<k. P n" shows "∀n<Suc k. P n" proof (intro allI impI) fix n assume "n < Suc k" show "P n" proof (cases "n < k") case True with Pn show ?thesis by simp next case False with `n < Suc k` have "n = k" by simp with Pk show ?thesis by simp qed qed lemma step_witness: assumes "in_cdesc RD1 y x" assumes "in_cdesc RD2 z y" shows "¬ no_step RD1 RD2" using assms by (cases RD1, cases RD2) (auto simp:no_step_def) theorem SCT_on_relations: assumes R: "R = mk_rel RDs" assumes sound: "sound_int \<A> RDs M" assumes "SCT \<A>" shows "∀x. accp R x" proof (rule, rule classical) fix x assume "¬ accp R x" with non_acc_has_idseq have "∃s. idseq R s x" . then obtain s where "idseq R s x" .. hence "∃cs. ∀i. cs i ∈ set RDs ∧ in_cdesc (cs i) (s (Suc i)) (s i)" unfolding R by (rule ex_cs) then obtain cs where [simp]: "!!i. cs i ∈ set RDs" and ird[simp]: "!!i. in_cdesc (cs i) (s (Suc i)) (s i)" by blast let ?cis = "λi. index_of RDs (cs i)" have "∀i. ∃G. (\<A> \<turnstile> ?cis i \<leadsto>G (?cis (Suc i))) ∧ approx G (RDs ! ?cis i) (RDs ! ?cis (Suc i)) (M ! ?cis i) (M ! ?cis (Suc i))" (is "∀i. ∃G. ?P i G") proof fix i let ?n = "?cis i" and ?n' = "?cis (Suc i)" have "in_cdesc (RDs ! ?n) (s (Suc i)) (s i)" "in_cdesc (RDs ! ?n') (s (Suc (Suc i))) (s (Suc i))" by (simp_all add:index_of_member) with step_witness have "¬ no_step (RDs ! ?n) (RDs ! ?n')" . moreover have "?n < length RDs" "?n' < length RDs" by (simp_all add:index_of_length[symmetric]) ultimately obtain G where "\<A> \<turnstile> ?n \<leadsto>G ?n'" and "approx G (RDs ! ?n) (RDs ! ?n') (M ! ?n) (M ! ?n')" using sound unfolding sound_int_def by auto thus "∃G. ?P i G" by blast qed with choice have "∃Gs. ∀i. ?P i (Gs i)" . then obtain Gs where A: "!!i. \<A> \<turnstile> ?cis i \<leadsto>(Gs i) (?cis (Suc i))" and B: "!!i. approx (Gs i) (RDs ! ?cis i) (RDs ! ?cis (Suc i)) (M ! ?cis i) (M ! ?cis (Suc i))" by blast let ?p = "λi. (?cis i, Gs i)" from A have "has_ipath \<A> ?p" unfolding has_ipath_def by auto with `SCT \<A>` SCT_def obtain th where "is_desc_thread th ?p" by auto then obtain n where fr: "∀i≥n. eqlat ?p th i" and inf: "∃∞i. descat ?p th i" unfolding is_desc_thread_def by auto from B have approx: "!!i. approx (Gs i) (cs i) (cs (Suc i)) (M ! ?cis i) (M ! ?cis (Suc i))" by (simp add:index_of_member) let ?seq = "λi. (M ! ?cis i) (th i) (s i)" have "!!i. n < i ==> ?seq (Suc i) ≤ ?seq i" proof - fix i let ?q1 = "th i" and ?q2 = "th (Suc i)" assume "n < i" with fr have "eqlat ?p th i" by simp hence "dsc (Gs i) ?q1 ?q2 ∨ eq (Gs i) ?q1 ?q2" by simp thus "?seq (Suc i) ≤ ?seq i" proof assume "dsc (Gs i) ?q1 ?q2" with approx have a:"decr (cs i) (cs (Suc i)) ((M ! ?cis i) ?q1) ((M ! ?cis (Suc i)) ?q2)" unfolding approx_def by auto show ?thesis apply (rule less_imp_le) apply (rule decr_in_cdesc[of _ "s (Suc i)" "s i"]) by (rule ird a)+ next assume "eq (Gs i) ?q1 ?q2" with approx have a:"decreq (cs i) (cs (Suc i)) ((M ! ?cis i) ?q1) ((M ! ?cis (Suc i)) ?q2)" unfolding approx_def by auto show ?thesis apply (rule decreq_in_cdesc[of _ "s (Suc i)" "s i"]) by (rule ird a)+ qed qed moreover have "∃∞i. ?seq (Suc i) < ?seq i" unfolding INF_nat proof fix i from inf obtain j where "i < j" and d: "descat ?p th j" unfolding INF_nat by auto let ?q1 = "th j" and ?q2 = "th (Suc j)" from d have "dsc (Gs j) ?q1 ?q2" by auto with approx have a:"decr (cs j) (cs (Suc j)) ((M ! ?cis j) ?q1) ((M ! ?cis (Suc j)) ?q2)" unfolding approx_def by auto have "?seq (Suc j) < ?seq j" apply (rule decr_in_cdesc[of _ "s (Suc j)" "s j"]) by (rule ird a)+ with `i < j` show "∃j. i < j ∧ ?seq (Suc j) < ?seq j" by auto qed ultimately have False by (rule no_inf_desc_nat_sequence[of "Suc n"]) simp thus "accp R x" .. qed end
lemma not_acc_smaller:
¬ accp R x ==> ∃y. R y x ∧ ¬ accp R y
lemma non_acc_has_idseq:
¬ accp R x ==> ∃s. idseq R s x
lemma some_rd:
mk_rel rds x y ==> ∃rd∈set rds. in_cdesc rd x y
lemma ex_cs:
idseq (mk_rel rds) s x
==> ∃cs. ∀i. cs i ∈ set rds ∧ in_cdesc (cs i) (s (Suc i)) (s i)
lemma decr_in_cdesc:
[| in_cdesc RD1.0 y x; in_cdesc RD2.0 z y; decr RD1.0 RD2.0 m1.0 m2.0 |]
==> m2.0 y < m1.0 x
lemma decreq_in_cdesc:
[| in_cdesc RD1.0 y x; in_cdesc RD2.0 z y; decreq RD1.0 RD2.0 m1.0 m2.0 |]
==> m2.0 y ≤ m1.0 x
lemma no_inf_desc_nat_sequence:
[| !!i. n ≤ i ==> s (Suc i) ≤ s i; ∃∞i. s (Suc i) < s i |] ==> False
lemma approx_empty:
approx (Graph {}) c1.0 c2.0 ms1.0 ms2.0
lemma approx_less:
[| stepP c1.0 c2.0 (ms1.0 i) (ms2.0 j) op <;
approx (Graph Es) c1.0 c2.0 ms1.0 ms2.0 |]
==> approx (Graph (insert (i, \<down>, j) Es)) c1.0 c2.0 ms1.0 ms2.0
lemma approx_leq:
[| stepP c1.0 c2.0 (ms1.0 i) (ms2.0 j) op ≤;
approx (Graph Es) c1.0 c2.0 ms1.0 ms2.0 |]
==> approx (Graph (insert (i, \<Down>, j) Es)) c1.0 c2.0 ms1.0 ms2.0
lemma no_stepI:
stepP c1.0 c2.0 m1.0 m2.0 (λx y. False) ==> no_step c1.0 c2.0
lemma length_simps:
length [] = 0
length (x # xs) = Suc (length xs)
lemma all_less_zero:
∀n<0. P n
lemma all_less_Suc:
[| P k; ∀n<k. P n |] ==> ∀n<Suc k. P n
lemma step_witness:
[| in_cdesc RD1.0 y x; in_cdesc RD2.0 z y |] ==> ¬ no_step RD1.0 RD2.0
theorem SCT_on_relations:
[| R = mk_rel RDs; sound_int \<A> RDs M; SCT \<A> |] ==> ∀x. accp R x