# HG changeset patch # User Cezary Kaliszyk # Date 1259229144 -3600 # Node ID 196aa25daadfb2fe6f21beb9321599e2c87056bf # Parent 98ccde1c184c4a93452ede3cbb7d060be5e825c6 Playing with Monos in LFex. diff -r 98ccde1c184c -r 196aa25daadf LFex.thy --- a/LFex.thy Thu Nov 26 10:32:31 2009 +0100 +++ b/LFex.thy Thu Nov 26 10:52:24 2009 +0100 @@ -197,6 +197,51 @@ thm akind_aty_atrm.induct +lemma "\P1 TYP TYP; \A A' K K' x. \(A::TY) = A'; P2 A A'; (K::KIND) = K'; P1 K K'\ \ P1 (KPI A x K) (KPI A' x K'); + \A A' K x x' K'. + \(A ::TY) = A'; P2 A A'; (K :: KIND) = ([(x, x')] \ K'); P1 K ([(x, x')] \ K'); x \ FV_ty A'; x \ FV_kind K' - {x'}\ + \ P1 (KPI A x K) (KPI A' x' K'); + \i j. i = j \ P2 (TCONST i) (TCONST j); + \A A' M M'. \(A ::TY) = A'; P2 A A'; (M :: TRM) = M'; P3 M M'\ \ P2 (TAPP A M) (TAPP A' M'); + \A A' B B' x. \(A ::TY) = A'; P2 A A'; (B ::TY) = B'; P2 B B'\ \ P2 (TPI A x B) (TPI A' x B'); + \A A' B x x' B'. + \(A ::TY) = A'; P2 A A'; (B ::TY) = ([(x, x')] \ B'); P2 B ([(x, x')] \ B'); x \ FV_ty B'; x \ FV_ty B' - {x'}\ + \ P2 (TPI A x B) (TPI A' x' B'); + \i j m. i = j \ P3 (CONS i) (m (CONS j)); \x y m. x = y \ P3 (VAR x) (m (VAR y)); + \M m M' N N'. \(M :: TRM) = m M'; P3 M (m M'); (N :: TRM) = N'; P3 N N'\ \ P3 (APP M N) (APP M' N'); + \A A' M M' x. \(A ::TY) = A'; P2 A A'; (M :: TRM) = M'; P3 M M'\ \ P3 (LAM A x M) (LAM A' x M'); + \A A' M x x' M' B'. + \(A ::TY) = A'; P2 A A'; (M :: TRM) = ([(x, x')] \ M'); P3 M ([(x, x')] \ M'); x \ FV_ty B'; x \ FV_trm M' - {x'}\ + \ P3 (LAM A x M) (LAM A' x' M')\ +\ ((x1 :: KIND) = x2 \ P1 x1 x2) \ + ((x3 ::TY) = x4 \ P2 x3 x4) \ ((x5 :: TRM) = x6 \ P3 x5 x6)" +apply(tactic {* procedure_tac @{context} @{thm akind_aty_atrm.induct} 1 *}) +apply(atomize (full)) +apply(rule my_equiv_res_forallR) +apply(tactic {* resolve_tac (Inductive.get_monos @{context}) 1 *}) +apply(rule my_equiv_res_forallR) +apply(tactic {* REPEAT_ALL_NEW (resolve_tac (Inductive.get_monos @{context})) 1 *}) +apply(rule my_equiv_res_forallR) +apply(tactic {* REPEAT_ALL_NEW (resolve_tac (Inductive.get_monos @{context})) 1 *}) +apply(rule my_equiv_res_forallR) +apply(tactic {* REPEAT_ALL_NEW (resolve_tac (Inductive.get_monos @{context})) 1 *}) +apply(rule my_equiv_res_forallR) +apply(tactic {* REPEAT_ALL_NEW (resolve_tac (Inductive.get_monos @{context})) 1 *}) +apply(rule my_equiv_res_forallR) +apply(tactic {* REPEAT_ALL_NEW (resolve_tac (Inductive.get_monos @{context})) 1 *}) +apply(rule my_equiv_res_forallR) +apply(tactic {* REPEAT_ALL_NEW (resolve_tac (Inductive.get_monos @{context})) 1 *}) +apply(rule my_equiv_res_forallR) +apply(tactic {* REPEAT_ALL_NEW (resolve_tac (Inductive.get_monos @{context})) 1 *}) +apply(rule my_equiv_res_forallR) +apply(tactic {* (resolve_tac (Inductive.get_monos @{context})) 1 *}) +apply(rule Set.imp_mono) +apply(rule impI) +apply(assumption) +apply(rule Set.imp_mono) + + + ML {* val rty_qty_rel = [(@{typ kind}, (@{typ KIND}, @{term akind})),