From: Wilmer Ricciotti Date: Sun, 27 Jan 2013 16:31:44 +0000 (+0000) Subject: well-foundedness done X-Git-Tag: make_still_working~1304 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=61a4954847fe3ab75f406573c14645cfd908a79e;p=helm.git well-foundedness done --- diff --git a/matita/matita/lib/turing/multi_universal/match.ma b/matita/matita/lib/turing/multi_universal/match.ma index 181a35235..a08ba7e04 100644 --- a/matita/matita/lib/turing/multi_universal/match.ma +++ b/matita/matita/lib/turing/multi_universal/match.ma @@ -701,63 +701,31 @@ lemma sem_match_step_termination : normalize #H destruct (H) ] ] ] qed. -axiom terminate_match_m : +(* lemma WF_to_WF_f : ∀A,B,R,f,b. WF A R (f b) → WF B (λx,y.R (f x) (f y)) b. *) +let rec WF_to_WF_f A B R f b (Hwf: WF A R (f b)) on Hwf: WF B (λx,y.R (f x) (f y)) b ≝ + match Hwf return (λa0,r.f b = a0 → WF B (λx,y:B. R (f x) (f y)) b) with + [ wf a Hwfa ⇒ λHeq.? ] (refl ??). +% #b1 #HRb @WF_to_WF_f @Hwfa nth_change_vec_neq [|@sym_not_eq //] - >Hmid_src #HR cases (HR ?? (refl ??)) -HR - >nth_change_vec // >Htape_dst #_ * #s0 * normalize in ⊢ (%→?); #H destruct (H) -| #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //] - >Hmid_src #HR cases (HR ?? (refl ??)) -HR - >nth_change_vec // >Htape_dst #_ normalize in ⊢ (%→?); - * #s0 * #H destruct (H) -| #x0 #xs0 #Htape_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //] - >Hmid_src #HR cases (HR ?? (refl ??)) -HR - >nth_change_vec // >Htape_dst #_ normalize in ⊢ (%→?); - * #s0 * #H destruct (H) -| #ls #s #rs lapply s -s lapply ls -ls lapply Hmid_src lapply t -t elim rs - [#t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?); >nth_change_vec_neq [|@sym_not_eq //] - >Hmid_src >nth_change_vec // >Hmid_dst #HR cases (HR ?? (refl ??)) -HR - >change_vec_change_vec #Ht1 #_ % #t2 whd in ⊢ (%→?); - >Ht1 >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src #HR - cases (HR ?? (refl ??)) -HR #_ - >nth_change_vec // * #s1 * normalize in ⊢ (%→?); #H destruct (H) - |#r0 #rs0 #IH #t #Hmid_src #ls #s #Hmid_dst % #t1 whd in ⊢ (%→?); - >nth_change_vec_neq [|@sym_not_eq //] >Hmid_src - #Htrue cases (Htrue ?? (refl ??)) -Htrue >change_vec_change_vec - >nth_change_vec // >Hmid_dst whd in match (tape_move_mono ???); #Ht1 - * #s0 * whd in ⊢ (??%?→?); #H destruct (H) #_ >Ht1 - lapply (IH t1 ? (s0::ls) r0 ?) - [ >Ht1 >nth_change_vec // - | >Ht1 >nth_change_vec_neq [|@sym_not_eq //] @Hmid_src - | >Ht1 >nth_change_vec // ] - ] -] -qed. *) +change with (λx,y.f x < f y) in ⊢ (??%?); @WF_to_WF_f @lt_WF +qed. lemma sem_match_m : ∀src,dst,sig,n. src ≠ dst → src < S n → dst < S n → match_m src dst sig n \vDash R_match_m src dst sig n. #src #dst #sig #n #Hneq #Hsrc #Hdst @WRealize_to_Realize [/2/| @wsem_match_m // ] -qed. +qed. \ No newline at end of file