preliminary tests
authorChristian Urban <urbanc@in.tum.de>
Wed, 14 Apr 2010 10:29:34 +0200
changeset 1828 f374ffd50c7c
parent 1827 eef70e8fa9ee
child 1829 ac8cb569a17b
preliminary tests
Nominal/Ex/Lambda.thy
--- a/Nominal/Ex/Lambda.thy	Wed Apr 14 10:28:17 2010 +0200
+++ b/Nominal/Ex/Lambda.thy	Wed Apr 14 10:29:34 2010 +0200
@@ -120,63 +120,132 @@
   "valid []"
 | "\<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) 
+where
+    t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
+  | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2 \<or> \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
+  | t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam x t : T1 \<rightarrow> T2"
+
+
 ML {*
-fun my_tac ctxt intros =  
- Nominal_Permeq.eqvt_strict_tac ctxt [] []
- THEN' resolve_tac intros 
- THEN_ALL_NEW 
-   (atac ORELSE'
-    EVERY'
-      [ rtac (Drule.instantiate' [] [SOME @{cterm "- p::perm"}] @{thm permute_boolE}),
-        Nominal_Permeq.eqvt_strict_tac ctxt @{thms permute_minus_cancel(2)} [],
-        atac ])
+fun map_term f t = 
+  (case f t of
+     NONE => map_term' f t 
+   | x => x)
+and map_term' f (t $ u) = 
+    (case (map_term f t, map_term f u) of
+        (NONE, NONE) => NONE
+      | (SOME t'', NONE) => SOME (t'' $ u)
+      | (NONE, SOME u'') => SOME (t $ u'')
+      | (SOME t'', SOME u'') => SOME (t'' $ u''))
+  | map_term' f (Abs (s, T, t)) = 
+      (case map_term f t of
+        NONE => NONE
+      | SOME t'' => SOME (Abs (s, T, t'')))
+  | map_term' _ _  = NONE;
+
+fun map_thm_tac ctxt tac thm =
+let
+  val monos = Inductive.get_monos ctxt
+in
+  EVERY [cut_facts_tac [thm] 1, etac rev_mp 1,
+    REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),
+    REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))]
+end
+
+(* 
+ proves F[f t] from F[t] where F[t] is the given theorem  
+  
+  - F needs to be monotone
+  - f returns either SOME for a term it fires 
+    and NONE elsewhere 
+*)
+fun map_thm ctxt f tac thm =
+let
+  val opt_goal_trm = map_term f (prop_of thm)
+  fun prove goal = 
+    Goal.prove ctxt [] [] goal (fn _ => map_thm_tac ctxt tac thm)
+in
+  case opt_goal_trm of
+    NONE => thm
+  | SOME goal => prove goal
+end
+
+fun transform_prem ctxt names thm =
+let
+  fun split_conj names (Const ("op &", _) $ p $ q) = 
+      (case head_of p of
+         Const (name, _) => if name mem names then SOME q else NONE
+       | _ => NONE)
+  | split_conj _ _ = NONE;
+in
+  map_thm ctxt (split_conj names) (etac conjunct2 1) thm
+end
 *}
 
+ML {*
+open Nominal_Permeq
+*}
+
+ML {* 
+fun single_case_tac ctxt pred_names pi intro  = 
+let
+  val rule = Drule.instantiate' [] [SOME pi] @{thm permute_boolE}
+in
+  eqvt_strict_tac ctxt [] [] THEN' 
+  SUBPROOF (fn {prems, context as ctxt, ...} =>
+    let
+      val prems' = map (transform_prem ctxt pred_names) prems
+      val side_cond_tac = EVERY' 
+        [ rtac rule, 
+          eqvt_strict_tac ctxt @{thms permute_minus_cancel(2)} [],
+          resolve_tac prems' ]
+    in
+      HEADGOAL (rtac intro THEN_ALL_NEW (resolve_tac prems' ORELSE' side_cond_tac)) 
+    end) ctxt
+end
+*}
+
+ML {*
+fun eqvt_rel_tac pred_name = 
+let
+  val thy = ProofContext.theory_of ctxt
+  val ({names, ...}, {raw_induct, intrs, ...}) =
+    Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
+  val param_no = length (Inductive.params_of raw_induct)
+  val (([raw_concl], [pi]), ctxt') = 
+    ctxt |> Variable.import_terms false [concl_of raw_induct] 
+         ||>> Variable.variant_fixes ["pi"];
+  val preds = map (fst o HOLogic.dest_imp)
+    (HOLogic.dest_conj (HOLogic.dest_Trueprop raw_concl));
+in
+
+end
+*}
+
+
+
 lemma [eqvt]:
   assumes a: "valid Gamma" 
   shows "valid (p \<bullet> Gamma)"
 using a
 apply(induct)
-apply(tactic {* my_tac @{context} @{thms valid.intros} 1 *})
-apply(tactic {* my_tac @{context} @{thms valid.intros} 1 *})
-done
-
-lemma
-  shows "valid Gamma \<longrightarrow> valid (p \<bullet> Gamma)"
-ML_prf {*
-val ({names, ...}, {raw_induct, intrs, elims, ...}) =
-      Inductive.the_inductive @{context} (Sign.intern_const @{theory} "valid")
-*}
-apply(tactic {* rtac raw_induct 1 *})
-apply(tactic {* my_tac @{context} @{thms valid.intros} 1 *})
-apply(tactic {* my_tac @{context} @{thms valid.intros} 1 *})
+apply(tactic {* my_tac @{context} ["Lambda.valid"] @{cterm "- p"} @{thm valid.intros(1)} 1 *})
+apply(tactic {* my_tac @{context }["Lambda.valid"] @{cterm "- p"} @{thm valid.intros(2)} 1 *})
 done
 
-
-thm eqvts
-thm eqvts_raw
-
-inductive
-  typing :: "(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60) 
-where
-    t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x, T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
-  | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1 \<rightarrow> T2 \<and> \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
-  | t_Lam[intro]: "\<lbrakk>atom x \<sharp> \<Gamma>; (x, T1) # \<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam x t : T1 \<rightarrow> T2"
-
-
-ML {* Inductive.the_inductive @{context} (Sign.intern_const @{theory} "typing") *}
-
 lemma 
   shows "Gamma \<turnstile> t : T \<longrightarrow> (p \<bullet> Gamma) \<turnstile> (p \<bullet> t) : (p \<bullet> T)"
 ML_prf {*
-val ({names, ...}, {raw_induct, intrs, elims, ...}) =
+val ({names, ...}, {raw_induct, ...}) =
       Inductive.the_inductive @{context} (Sign.intern_const @{theory} "typing")
 *}
 apply(tactic {* rtac raw_induct 1 *})
-apply(tactic {* my_tac @{context} @{thms typing.intros} 1 *})
-apply(perm_strict_simp)
-apply(rule typing.intros)
-oops
+apply(tactic {* my_tac @{context} ["Lambda.typing"] @{cterm "- p"} @{thm typing.intros(1)} 1 *})
+apply(tactic {* my_tac @{context} ["Lambda.typing"] @{cterm "- p"} @{thm typing.intros(2)} 1 *})
+apply(tactic {* my_tac @{context} ["Lambda.typing"] @{cterm "- p"} @{thm typing.intros(3)} 1 *})
+done
 
 lemma uu[eqvt]:
   assumes a: "Gamma \<turnstile> t : T"