strong rule inductions; as an example the weakening lemma works
authorChristian Urban <urbanc@in.tum.de>
Wed, 05 Jan 2011 16:51:27 +0000
changeset 2638 e1e2ca92760b
parent 2637 3890483c674f
child 2639 a8fc346deda3
strong rule inductions; as an example the weakening lemma works
Nominal/Ex/Typing.thy
--- a/Nominal/Ex/Typing.thy	Tue Jan 04 13:47:38 2011 +0000
+++ b/Nominal/Ex/Typing.thy	Wed Jan 05 16:51:27 2011 +0000
@@ -21,8 +21,11 @@
 thm lam.size_eqvt
 
 ML {*
+fun mk_cplus p q = Thm.capply (Thm.capply @{cterm "plus::perm \<Rightarrow> perm \<Rightarrow> perm"} p) q 
+
 fun mk_cminus p = Thm.capply @{cterm "uminus::perm \<Rightarrow> perm"} p 
 
+
 fun minus_permute_intro_tac p = 
   rtac (Drule.instantiate' [] [SOME (mk_cminus p)] @{thm permute_boolE})
 
@@ -48,8 +51,13 @@
       |> HOLogic.mk_Trueprop
       |> (curry Logic.list_implies) prems
       |> (curry list_all_free) params
+    val finite_goal = avoid_trm
+      |> mk_finite
+      |> HOLogic.mk_Trueprop
+      |> (curry Logic.list_implies) prems
+      |> (curry list_all_free) params
   in 
-    if null avoid then [] else [vc_goal]
+    if null avoid then [] else [vc_goal, finite_goal]
   end
 *}
 
@@ -136,6 +144,12 @@
 *}
 
 ML {*
+fun map7 _ [] [] [] [] [] [] [] = []
+  | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (r :: rs) (s :: ss) = 
+      f x y z u v r s :: map7 f xs ys zs us vs rs ss
+*}
+
+ML {*
 (* local abbreviations *)
 fun eqvt_stac ctxt = Nominal_Permeq.eqvt_strict_tac ctxt @{thms permute_minus_cancel} []  
 fun eqvt_srule ctxt = Nominal_Permeq.eqvt_strict_rule ctxt @{thms permute_minus_cancel} []  
@@ -157,12 +171,16 @@
       val prems' = prems
         |> map (minus_permute_elim p)
         |> map (eqvt_srule context)
-     
+
       val prm' = (prems' MRS prm)
         |> flag ? (all_elims [p])
-        |> flag ? (simplify (HOL_basic_ss addsimps @{thms permute_minus_cancel}))
+        |> flag ? (eqvt_srule context)
+
+      val _ = tracing ("prm':" ^ @{make_string} prm')
     in
-      simp_tac (HOL_ss addsimps (prm' :: @{thms induct_forall_def })) 1
+      print_tac "start helper"
+      THEN asm_full_simp_tac (HOL_ss addsimps (prm' :: @{thms induct_forall_def})) 1
+      THEN print_tac "final helper"
     end) ctxt
 *}
 
@@ -178,7 +196,7 @@
       val prem' = cterm_instantiate (intr_cvars ~~ p_prms) prem
 
       (* for inductive-premises*)
-      fun tac1 prm = helper_tac true prm p context
+      fun tac1 prm = helper_tac true prm p context 
 
       (* for non-inductive premises *)   
       fun tac2 prm =  
@@ -189,63 +207,120 @@
       fun select prm (t, i) =
         (if member same_name Ps (real_head_of t) then tac1 prm else tac2 prm) i
     in
-      EVERY1 [rtac prem', RANGE (map (SUBGOAL o select) prems) ]
+      EVERY1 [eqvt_stac ctxt, rtac prem', RANGE (map (SUBGOAL o select) prems) ]
     end) ctxt
 *}
 
 
 ML {*
-fun fresh_thm ctxt fresh_thms p c prms avoid_trm =
+fun fresh_thm ctxt user_thm p c concl_args avoid_trm =
   let
     val conj1 = 
       mk_fresh_star (mk_perm (Bound 0) (mk_perm p avoid_trm)) c
     val conj2 =
-      mk_fresh_star_ty @{typ perm} (mk_supp (HOLogic.mk_tuple (map (mk_perm p) prms))) (Bound 0)
+      mk_fresh_star_ty @{typ perm} (mk_supp (HOLogic.mk_tuple (map (mk_perm p) concl_args))) (Bound 0)
     val fresh_goal = mk_exists ("q", @{typ perm}) (HOLogic.mk_conj (conj1, conj2))
       |> HOLogic.mk_Trueprop
 
-    val _ = tracing ("fresh goal: " ^ Syntax.string_of_term ctxt fresh_goal)
+    val ss = @{thms finite_supp supp_Pair finite_Un permute_finite} @ 
+             @{thms fresh_star_Pair fresh_star_permute_iff}
+    val simp = asm_full_simp_tac (HOL_ss addsimps ss)
   in 
     Goal.prove ctxt [] [] fresh_goal
-      (K (HEADGOAL (rtac @{thm at_set_avoiding2})))
+      (K (HEADGOAL (rtac @{thm at_set_avoiding2} 
+          THEN_ALL_NEW EVERY' [cut_facts_tac user_thm, REPEAT o etac @{thm conjE}, simp])))
+  end
+*}
+
+ML {* 
+val supp_perm_eq' = 
+  @{lemma "supp (p \<bullet> x) \<sharp>* q ==> p \<bullet> x == (q + p) \<bullet> x" by (simp add: supp_perm_eq)}
+val fresh_star_plus =
+  @{lemma "(q \<bullet> (p \<bullet> x)) \<sharp>* c ==> ((q + p) \<bullet> x) \<sharp>* c" by (simp add: permute_plus)}
+*}
+
+ML {*
+fun binder_tac prem intr_cvars param_trms Ps user_thm avoid avoid_trm concl_args ctxt = 
+  Subgoal.FOCUS (fn {context = ctxt, params, prems, concl, ...} =>
+    let
+      val thy = ProofContext.theory_of ctxt
+      val (prms, p, c) = split_last2 (map snd params)
+      val prm_trms = map term_of prms
+      val prm_tys = map fastype_of prm_trms
+
+      val avoid_trm' = subst_free (param_trms ~~ prm_trms) avoid_trm 
+      val concl_args' = map (subst_free (param_trms ~~ prm_trms)) concl_args 
+      
+      val user_thm' = map (cterm_instantiate (intr_cvars ~~ prms)) user_thm
+        |> map (full_simplify (HOL_ss addsimps (@{thm fresh_star_Pair}::prems)))
+      
+      val fthm = fresh_thm ctxt user_thm' (term_of p) (term_of c) concl_args' avoid_trm'
+
+      val (([(_, q)], fprop :: fresh_eqs), ctxt') = Obtain.result
+              (K (EVERY1 [etac @{thm exE}, 
+                          full_simp_tac (HOL_basic_ss addsimps @{thms supp_Pair fresh_star_Un}), 
+                          REPEAT o etac @{thm conjE},
+                          dtac fresh_star_plus,
+                          REPEAT o dtac supp_perm_eq'])) [fthm] ctxt 
+
+      val expand_conv = Conv.try_conv (Conv.rewrs_conv fresh_eqs)
+      fun expand_conv_bot ctxt = Conv.bottom_conv (K expand_conv) ctxt
+
+      val cperms = map (cterm_of thy o perm_const) prm_tys
+      val qp_prms = map2 (fn ct1 => fn ct2 => Thm.mk_binop ct1 (mk_cplus q p) ct2) cperms prms 
+      val prem' = cterm_instantiate (intr_cvars ~~ qp_prms) prem
+
+      val fprop' = eqvt_srule ctxt' fprop 
+      val tac_fresh = simp_tac (HOL_basic_ss addsimps [fprop'])
+
+      (* for inductive-premises*)
+      fun tac1 prm = helper_tac true prm (mk_cplus q p) ctxt' 
+
+      (* for non-inductive premises *)   
+      fun tac2 prm =  
+        EVERY' [ minus_permute_intro_tac (mk_cplus q p), 
+                 eqvt_stac ctxt, 
+                 helper_tac false prm (mk_cplus q p) ctxt' ]
+
+      fun select prm (t, i) =
+        (if member same_name Ps (real_head_of t) then tac1 prm else tac2 prm) i
+
+      val _ = tracing ("fthm:\n" ^ @{make_string} fthm)
+      val _ = tracing ("fr_eqs:\n" ^ cat_lines (map @{make_string} fresh_eqs))
+      val _ = tracing ("fprop:\n" ^ @{make_string} fprop)
+      val _ = tracing ("fprop':\n" ^ @{make_string} fprop')
+      val _ = tracing ("fperm:\n" ^ @{make_string} q)
+      val _ = tracing ("prem':\n" ^ @{make_string} prem')
+
+      val side_thm = Goal.prove ctxt' [] [] (term_of concl)
+        (fn {context, ...} => 
+           EVERY1 [ CONVERSION (expand_conv_bot context),
+                    eqvt_stac context,
+                    rtac prem',
+                    RANGE (tac_fresh :: map (SUBGOAL o select) prems),
+                    K (print_tac "GOAL") ])
+        |> singleton (ProofContext.export ctxt' ctxt)        
+    in
+      rtac side_thm 1
+    end) ctxt
+*}
+
+ML {*
+fun case_tac ctxt Ps avoid avoid_trm intr_cvars param_trms prem user_thm concl_args =
+  let
+    val tac1 = non_binder_tac prem intr_cvars Ps ctxt
+    val tac2 = binder_tac prem intr_cvars param_trms Ps user_thm avoid avoid_trm concl_args ctxt
+  in 
+    EVERY' [ rtac @{thm allI}, rtac @{thm allI}, if null avoid then tac1 else tac2 ]
   end
 *}
 
 ML {*
-fun binder_tac prem intr_cvars param_trms Ps fresh_thms avoid avoid_trm ctxt = 
-  Subgoal.FOCUS (fn {context, params, ...} =>
-    let
-      val thy = ProofContext.theory_of context
-      val (prms, p, c) = split_last2 (map snd params)
-      val prm_trms = map term_of prms
-      val prm_tys = map fastype_of prm_trms
-      val cperms = map (cterm_of thy o perm_const) prm_tys
-      val p_prms = map2 (fn ct1 => fn ct2 => Thm.mk_binop ct1 p ct2) cperms prms 
-      val prem' = cterm_instantiate (intr_cvars ~~ p_prms) prem
-      val avoid_trm' = subst_free (param_trms ~~ prm_trms) avoid_trm      
-
-      val fthm = fresh_thm context fresh_thms (term_of p) (term_of c) (map term_of prms) avoid_trm'
-    in
-      Skip_Proof.cheat_tac thy
-      (* EVERY1 [rtac prem'] *)  
-    end) ctxt
-*}
-
-ML {*
-fun case_tac ctxt fresh_thms Ps (avoid, avoid_trm) intr_cvars param_trms prem =
+fun prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms concl_args 
+  {prems, context} =
   let
-    val tac1 = non_binder_tac prem intr_cvars Ps ctxt
-    val tac2 = binder_tac prem intr_cvars param_trms Ps fresh_thms avoid avoid_trm ctxt
-  in 
-    EVERY' [ rtac @{thm allI}, rtac @{thm allI}, eqvt_stac ctxt,
-             if null avoid then tac1 else tac2 ]
-  end
-*}
-
-ML {*
-fun prove_sinduct_tac raw_induct fresh_thms Ps avoids avoid_trms intr_cvars param_trms {prems, context} =
-  let
-    val cases_tac = map4 (case_tac context fresh_thms Ps) (avoids ~~avoid_trms) intr_cvars param_trms prems
+    val cases_tac = 
+      map7 (case_tac context Ps) avoids avoid_trms intr_cvars param_trms prems user_thms concl_args
   in 
     EVERY1 [ DETERM o rtac raw_induct, RANGE cases_tac ]
   end
@@ -255,8 +330,10 @@
 val normalise = @{lemma "(Q --> (!p c. P p c)) ==> (!!c. Q ==> P (0::perm) c)" by simp}
 *}
 
+ML {* Local_Theory.note *}
+
 ML {*
-fun prove_strong_inductive rule_names avoids raw_induct intrs ctxt =
+fun prove_strong_inductive pred_names rule_names avoids raw_induct intrs ctxt =
   let
     val thy = ProofContext.theory_of ctxt
     val ((_, [raw_induct']), ctxt') = Variable.import true [raw_induct] ctxt
@@ -311,10 +388,10 @@
     val ind_prems' = ind_prems
       |> map2 (prep_prem Ps c_name c_ty) (avoids ~~ avoid_trms)   
 
-    fun after_qed ctxt_outside fresh_thms ctxt = 
+    fun after_qed ctxt_outside user_thms ctxt = 
       let
-        val thms = Goal.prove ctxt [] ind_prems' ind_concl' 
-          (prove_sinduct_tac raw_induct fresh_thms Ps avoids avoid_trms intr_cvars param_trms) 
+        val strong_ind_thms = Goal.prove ctxt [] ind_prems' ind_concl' 
+        (prove_sinduct_tac raw_induct user_thms Ps avoids avoid_trms intr_cvars param_trms intr_concls_args) 
           |> singleton (ProofContext.export ctxt ctxt_outside)
           |> Datatype_Aux.split_conj_thm
           |> map (fn thm => thm RS normalise)
@@ -322,10 +399,22 @@
           |> map (Drule.rotate_prems (length ind_prems'))
           |> map zero_var_indexes
 
-        val _ = tracing ("RESULTS\n" ^ cat_lines (map (Syntax.string_of_term ctxt o prop_of) thms))
+        val qualified_thm_name = pred_names
+          |> map Long_Name.base_name
+          |> space_implode "_"
+          |> (fn s => Binding.qualify false s (Binding.name "strong_induct"))
+
+        val attrs = 
+          [ Attrib.internal (K (Rule_Cases.consumes 1)),
+            Attrib.internal (K (Rule_Cases.case_names rule_names)) ]
+        val _ = tracing ("RESULTS\n" ^ cat_lines (map (Syntax.string_of_term ctxt o prop_of) strong_ind_thms))
+        val _ = tracing ("rule_names: " ^ commas rule_names)
+        val _ = tracing ("pred_names: " ^ commas pred_names)
       in
         ctxt
-      end 
+        |> Local_Theory.note ((qualified_thm_name, attrs), strong_ind_thms)    
+        |> snd   
+      end
   in
     Proof.theorem NONE (after_qed ctxt) ((map o map) (rpair []) vc_compat_goals) ctxt''
   end
@@ -366,7 +455,7 @@
 
     val avoid_trms = map2 read_avoids avoids_ordered intrs
   in
-    prove_strong_inductive rule_names avoid_trms raw_induct intrs ctxt
+    prove_strong_inductive names rule_names avoid_trms raw_induct intrs ctxt
   end
 *}
 
@@ -422,6 +511,8 @@
 
 nominal_inductive Acc .
 
+thm Acc.strong_induct
+
 section {* Typing *}
 
 nominal_datatype ty =
@@ -441,8 +532,8 @@
 inductive
   valid :: "(name \<times> ty) list \<Rightarrow> bool"
 where
-  "valid []"
-| "\<lbrakk>atom x \<sharp> Gamma; valid Gamma\<rbrakk> \<Longrightarrow> valid ((x, T)#Gamma)"
+  v_Nil[intro]: "valid []"
+| v_Cons[intro]: "\<lbrakk>atom x \<sharp> Gamma; valid Gamma\<rbrakk> \<Longrightarrow> valid ((x, T)#Gamma)"
 
 inductive
   typing :: "(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
@@ -459,105 +550,66 @@
 equivariance valid
 equivariance typing
 
-
 nominal_inductive typing
   avoids t_Lam: "x"
-      (* | t_Var: "x" *)
   apply -
   apply(simp_all add: fresh_star_def ty_fresh lam.fresh)?
   done
 
+thm typing.strong_induct
 
+abbreviation
+  "sub_context" :: "(name \<times> ty) list \<Rightarrow> (name \<times> ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60) 
+where
+  "\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>x T. (x, T) \<in> set \<Gamma>1 \<longrightarrow> (x, T) \<in> set \<Gamma>2"
+
+text {* Now it comes: The Weakening Lemma *}
+
+text {*
+  The first version is, after setting up the induction, 
+  completely automatic except for use of atomize. *}
 
-lemma
-  fixes c::"'a::fs"
-  assumes a: "\<Gamma> \<turnstile> t : T" 
-  and a1: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P c \<Gamma> (Var x) T"
-  and a2: "\<And>\<Gamma> t1 T1 T2 t2 c. \<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2; \<And>d. P d \<Gamma> t1 T1 \<rightarrow> T2; \<Gamma> \<turnstile> t2 : T1; \<And>d. P d \<Gamma> t2 T1\<rbrakk> 
-           \<Longrightarrow> P c \<Gamma> (App t1 t2) T2"
-  and a3: "\<And>x \<Gamma> T1 t T2 c. \<lbrakk>{atom x} \<sharp>* c; atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2; \<And>d. P d ((x, T1) # \<Gamma>) t T2\<rbrakk> 
-           \<Longrightarrow> P c \<Gamma> (Lam x t) T1 \<rightarrow> T2"
-  shows "P c \<Gamma> t T"
-proof -
-  from a have "\<And>p c. P c (p \<bullet> \<Gamma>) (p \<bullet> t) (p \<bullet> T)"
-  proof (induct)
-    case (t_Var \<Gamma> x T p c)
-    then show ?case
-      apply -
-      apply(perm_strict_simp)
-      thm a1
-      apply(rule a1)
-      apply(drule_tac p="p" in permute_boolI)
-      apply(perm_strict_simp add: permute_minus_cancel)
-      apply(assumption)
-      apply(rotate_tac 1)
-      apply(drule_tac p="p" in permute_boolI)
-      apply(perm_strict_simp add: permute_minus_cancel)
-      apply(assumption)
-      done
-  next
-    case (t_App \<Gamma> t1 T1 T2 t2 p c)
-    then show ?case
-      apply -
-      apply(perm_strict_simp)
-      apply(rule a2)
-      apply(drule_tac p="p" in permute_boolI)
-      apply(perm_strict_simp add: permute_minus_cancel)
-      apply(assumption)
-      apply(assumption)
-      apply(rotate_tac 2)
-      apply(drule_tac p="p" in permute_boolI)
-      apply(perm_strict_simp add: permute_minus_cancel)
-      apply(assumption)
-      apply(assumption)
-      done
-  next
-    case (t_Lam x \<Gamma> T1 t T2 p c)
-    then show ?case
-      apply -
-      apply(subgoal_tac "\<exists>q. (q \<bullet> {p \<bullet> atom x}) \<sharp>* c \<and> 
-        supp (p \<bullet> \<Gamma>, p \<bullet> Lam x t, p \<bullet> (T1 \<rightarrow> T2)) \<sharp>* q")
-      apply(erule exE)
-      apply(rule_tac t="p \<bullet> \<Gamma>" and  s="(q + p) \<bullet> \<Gamma>" in subst)
-      apply(simp only: permute_plus)
-      apply(rule supp_perm_eq)
-      apply(simp add: supp_Pair fresh_star_Un)
-      apply(rule_tac t="p \<bullet> Lam x t" and  s="(q + p) \<bullet> Lam x t" in subst)
-      apply(simp only: permute_plus)
-      apply(rule supp_perm_eq)
-      apply(simp add: supp_Pair fresh_star_Un)
-      apply(rule_tac t="p \<bullet> (T1 \<rightarrow> T2)" and  s="(q + p) \<bullet> (T1 \<rightarrow> T2)" in subst)
-      apply(simp only: permute_plus)
-      apply(rule supp_perm_eq)
-      apply(simp add: supp_Pair fresh_star_Un)
-      (* apply(perm_simp) *) 
-      apply(simp (no_asm) only: eqvts)
-      apply(rule a3)
-      apply(simp only: eqvts permute_plus)
-      apply(rule_tac p="- (q + p)" in permute_boolE)
-      apply(perm_strict_simp add: permute_minus_cancel)
-      apply(assumption)
-       apply(rule_tac p="- (q + p)" in permute_boolE)
-      apply(perm_strict_simp add: permute_minus_cancel)
-      apply(assumption)
-      apply(perm_strict_simp)
-      apply(simp only:)
-      thm at_set_avoiding2
-      apply(rule at_set_avoiding2)
-      apply(simp add: finite_supp)
-      apply(simp add: finite_supp)
-      apply(simp add: finite_supp)
-      apply(rule_tac p="-p" in permute_boolE)
-      apply(perm_strict_simp add: permute_minus_cancel)
-	--"supplied by the user"
-      apply(simp add: fresh_star_Pair)
-      sorry
-  qed
-  then have "P c (0 \<bullet> \<Gamma>) (0 \<bullet> t) (0 \<bullet> T)" .
-  then show "P c \<Gamma> t T" by simp
-qed
+lemma weakening_version2: 
+  fixes \<Gamma>1 \<Gamma>2::"(name \<times> ty) list"
+  and   t ::"lam"
+  and   \<tau> ::"ty"
+  assumes a: "\<Gamma>1 \<turnstile> t : T"
+  and     b: "valid \<Gamma>2" 
+  and     c: "\<Gamma>1 \<subseteq> \<Gamma>2"
+  shows "\<Gamma>2 \<turnstile> t : T"
+using a b c
+proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)
+  case (t_Var \<Gamma>1 x T)  (* variable case *)
+  have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact 
+  moreover  
+  have "valid \<Gamma>2" by fact 
+  moreover 
+  have "(x,T)\<in> set \<Gamma>1" by fact
+  ultimately show "\<Gamma>2 \<turnstile> Var x : T" by auto
+next
+  case (t_Lam x \<Gamma>1 T1 t T2) (* lambda case *)
+  have vc: "atom x \<sharp> \<Gamma>2" by fact   (* variable convention *)
+  have ih: "\<lbrakk>valid ((x, T1) # \<Gamma>2); (x, T1) # \<Gamma>1 \<subseteq> (x, T1) # \<Gamma>2\<rbrakk> \<Longrightarrow> (x, T1) # \<Gamma>2 \<turnstile> t : T2" by fact
+  have "\<Gamma>1 \<subseteq> \<Gamma>2" by fact
+  then have "(x, T1) # \<Gamma>1 \<subseteq> (x, T1) # \<Gamma>2" by simp
+  moreover
+  have "valid \<Gamma>2" by fact
+  then have "valid ((x, T1) # \<Gamma>2)" using vc by (simp add: v_Cons)
+  ultimately have "(x, T1) # \<Gamma>2 \<turnstile> t : T2" using ih by simp
+  with vc show "\<Gamma>2 \<turnstile> Lam x t : T1 \<rightarrow> T2" by auto
+qed (auto) (* app case *)
 
-*)
+lemma weakening_version1: 
+  fixes \<Gamma>1 \<Gamma>2::"(name \<times> ty) list"
+  assumes a: "\<Gamma>1 \<turnstile> t : T" 
+  and     b: "valid \<Gamma>2" 
+  and     c: "\<Gamma>1 \<subseteq> \<Gamma>2"
+  shows "\<Gamma>2 \<turnstile> t : T"
+using a b c
+apply (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)
+apply (auto | atomize)+
+done
+
 
 
 end