Nominal/Ex/Lambda.thy
changeset 1833 2050b5723c04
parent 1831 16653e702d89
child 1835 636de31888a6
--- a/Nominal/Ex/Lambda.thy	Wed Apr 14 13:21:38 2010 +0200
+++ b/Nominal/Ex/Lambda.thy	Wed Apr 14 14:41:54 2010 +0200
@@ -127,182 +127,31 @@
   | 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 {*
-val inductive_atomize = @{thms induct_atomize};
-
-val atomize_conv = 
-  MetaSimplifier.rewrite_cterm (true, false, false) (K (K NONE))
-    (HOL_basic_ss addsimps inductive_atomize);
-val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv);
-fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1
-  (Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt));
-*}
-
-ML {*
-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
-open Nominal_ThmDecls
-*}
-
-ML {*
-fun mk_perm p trm =
-let
-  val ty = fastype_of trm
-in
-  Const (@{const_name "permute"}, @{typ "perm"} --> ty --> ty) $ p $ trm
-end
-
-fun mk_minus p = 
- Const (@{const_name "uminus"}, @{typ "perm => perm"}) $ p   
-*}
+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 {* 
-fun single_case_tac ctxt pred_names pi intro  = 
-let
-  val thy = ProofContext.theory_of ctxt
-  val cpi = Thm.cterm_of thy (mk_minus pi)
-  val rule = Drule.instantiate' [] [SOME cpi] @{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 prepare_pred params_no pi pred =
-let
-  val (c, xs) = strip_comb pred;
-  val (xs1, xs2) = chop params_no xs
-in
-  HOLogic.mk_imp 
-    (pred, list_comb (c, xs1 @ map (mk_perm pi) xs2))
-end
-*}
-
-ML {*
-fun transp ([] :: _) = []
-  | transp xs = map hd xs :: transp (map tl xs);
-*}
-
-ML {* 
-  Local_Theory.note;
-  Local_Theory.notes;
-  fold_map
-*}
+inductive
+  typing2' :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ 2\<Turnstile> _ : _" [60,60,60] 60) 
+where
+    t2'_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> 2\<Turnstile> Var x : T"
+  | t2'_App[intro]: "\<lbrakk>\<Gamma> 2\<Turnstile> t1 : T1\<rightarrow>T2 \<or> \<Gamma> 2\<Turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> 2\<Turnstile> App t1 t2 : T2"
+  | t2'_Lam[intro]: "\<lbrakk>atom x\<sharp>\<Gamma>;(x,T1)#\<Gamma> 2\<Turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> 2\<Turnstile> Lam x t : T1\<rightarrow>T2"
 
-ML {*
-fun note_named_thm (name, thm) ctxt = 
-let
-  val thm_name = Binding.qualified_name 
-    (Long_Name.qualify (Long_Name.base_name name) "eqvt")
-  val attr = Attrib.internal (K eqvt_add)
-in
-  Local_Theory.note ((thm_name, [attr]), [thm]) ctxt
-end
-*}
+inductive
+  typing'' :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ |\<Turnstile> _ : _" [60,60,60] 60)
+and  valid' :: "(name\<times>ty) list \<Rightarrow> bool"
+where
+    v1[intro]: "valid' []"
+  | v2[intro]: "\<lbrakk>valid' \<Gamma>;atom x\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid' ((x,T)#\<Gamma>)"
+  | 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; \<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 eqvt_rel_tac pred_name ctxt = 
-let
-  val thy = ProofContext.theory_of ctxt
-  val ({names, ...}, {raw_induct, intrs, ...}) =
-    Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
-  val raw_induct = atomize_induct ctxt raw_induct;
-  val intros = map atomize_intr intrs;
-  val params_no = length (Inductive.params_of raw_induct)
-  val (([raw_concl], [raw_pi]), ctxt') = 
-    ctxt |> Variable.import_terms false [concl_of raw_induct] 
-         ||>> Variable.variant_fixes ["pi"]
-  val pi = Free (raw_pi, @{typ perm})
-  val preds = map (fst o HOLogic.dest_imp)
-    (HOLogic.dest_conj (HOLogic.dest_Trueprop raw_concl));
-  val goal = HOLogic.mk_Trueprop 
-    (foldr1 HOLogic.mk_conj (map (prepare_pred params_no pi) preds))
-  val thm = Goal.prove ctxt' [] [] goal (fn {context,...} => 
-    HEADGOAL (EVERY' (rtac raw_induct :: map (single_case_tac context names pi) intros)))
-    |> singleton (ProofContext.export ctxt' ctxt)
-  val thms = map (fn th => zero_var_indexes (th RS mp)) (Datatype_Aux.split_conj_thm thm)
-in
-   ctxt |> fold_map note_named_thm (names ~~ thms)
-        |> snd
-end
-*}
-
-
-ML {*
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ =
-  OuterSyntax.local_theory "equivariance"
-    "prove equivariance for inductive predicate involving nominal datatypes" K.thy_decl
-    (P.xname >> eqvt_rel_tac);
-
-end;
-*}
+use "../../Nominal-General/nominal_eqvt.ML"
 
 equivariance valid
 equivariance typing
@@ -312,6 +161,26 @@
 thm eqvts
 thm eqvts_raw
 
+equivariance typing'
+equivariance typing2' 
+equivariance typing''
+
+ML {*
+fun mk_minus p = 
+ Const (@{const_name "uminus"}, @{typ "perm => perm"}) $ p
+*}
+
+inductive
+  tt :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool)"  
+  for r :: "('a \<Rightarrow> 'a \<Rightarrow> bool)"
+where
+    "tt r a b"
+  | "tt r a b ==> tt r a c"
+
+(*
+equivariance tt
+*)
+
 
 inductive
  alpha_lam_raw'