Theory AOT_RestrictedVariables

(*<*)
theory AOT_RestrictedVariables
  imports AOT_PLM
  keywords "AOT_register_rigid_restricted_type" :: thy_goal
       and "AOT_register_restricted_type" :: thy_goal
begin
(*>*)

section‹Restricted Variables›

locale AOT_restriction_condition =
  fixes ψ :: 'a::AOT_Term_id_2  𝗈
  assumes "res-var:2"[AOT]: [v  α ψ{α}]
  assumes "res-var:3"[AOT]: [v  ψ{τ}  τ]

MLfun register_restricted_type (name:string, restriction:string) thy =
let
val ctxt = thy
val ctxt = setupStrictWorld ctxt
val trm = Syntax.check_term ctxt (AOT_read_term @{nonterminal φ'} ctxt restriction)
val free = case (Term.add_frees trm []) of [f] => f |
    _ => raise Term.TERM ("Expected single free variable.", [trm])
val trm = Term.absfree free trm
val localeTerm = Const (const_nameAOT_restriction_condition, dummyT) $ trm
val localeTerm = Syntax.check_term ctxt localeTerm
fun after_qed thms thy = let
val st = Interpretation.global_interpretation
  (([(@{locale AOT_restriction_condition}, ((name, true),
           (Expression.Named [("ψ", trm)], [])))], [])) [] thy
val st = Proof.refine_insert (flat thms) st
val thy = Proof.global_immediate_proof st

val thy = Local_Theory.background_theory
  (AOT_Constraints.map (Symtab.update
    (name, (term_of (snd free), term_of (snd free))))) thy
val thy = Local_Theory.background_theory
  (AOT_Restriction.map (Symtab.update
    (name, (trm, Const (const_nameAOT_term_of_var, dummyT))))) thy

in thy end
in
Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop localeTerm, [])]] ctxt
end

val _ =
  Outer_Syntax.command
    command_keywordAOT_register_restricted_type
    "Register a restricted type."
    (((Parse.short_ident --| Parse.$$$ ":") -- Parse.term) >>
    (Toplevel.local_theory_to_proof NONE NONE o register_restricted_type));

locale AOT_rigid_restriction_condition = AOT_restriction_condition +
  assumes rigid[AOT]: [v  α(ψ{α}  ψ{α})]
begin
lemma rigid_condition[AOT]: [v  (ψ{α}  ψ{α})]
  using rigid[THEN "∀E"(2)] RN by simp
lemma type_set_nonempty[AOT_no_atp, no_atp]: x . x  { α . [w0  ψ{α}]}
  by (metis "instantiation" mem_Collect_eq "res-var:2")
end

locale AOT_restricted_type = AOT_rigid_restriction_condition +
  fixes Rep and Abs
  assumes AOT_restricted_type_definition[AOT_no_atp]:
    type_definition Rep Abs { α . [w0  ψ{α}]}
begin

AOT_theorem restricted_var_condition: ψ{«AOT_term_of_var (Rep α)»}
proof -
  interpret type_definition Rep Abs "{ α . [w0  ψ{α}]}"
    using AOT_restricted_type_definition.
  AOT_actually {
    AOT_have «AOT_term_of_var (Rep α)» and ψ{«AOT_term_of_var (Rep α)»}
      using AOT_sem_imp Rep "res-var:3" by auto
  }
  moreover AOT_actually {
    AOT_have ψ{α}  ψ{α} for α
      using AOT_sem_box rigid_condition by presburger
    AOT_hence ψ{τ}  ψ{τ} if τ for τ
      by (metis AOT_model.AOT_term_of_var_cases AOT_sem_denotes that)
  }
  ultimately AOT_show ψ{«AOT_term_of_var (Rep α)»}
    using AOT_sem_box AOT_sem_imp by blast
qed
lemmas "ψ" = restricted_var_condition

AOT_theorem GEN: assumes for arbitrary α: φ{«AOT_term_of_var (Rep α)»}
  shows α (ψ{α}  φ{α})
proof(rule GEN; rule "→I")
  interpret type_definition Rep Abs "{ α . [w0  ψ{α}]}"
    using AOT_restricted_type_definition.
  fix α
  AOT_assume ψ{α}
  AOT_hence 𝒜ψ{α}
    by (metis AOT_model_axiom_def AOT_sem_box AOT_sem_imp act_closure rigid_condition)
  hence 0: [w0  ψ{α}] by (metis AOT_sem_act)
  {
    fix τ
    assume α_def: α = Rep τ
    AOT_have φ{α}
      unfolding α_def
      using assms by blast
  }
  AOT_thus φ{α}
    using Rep_cases[simplified, OF 0]
    by blast
qed
lemmas "∀I" = GEN

end


lemma AOT_restricted_type_intro[AOT_no_atp, no_atp]:
  assumes type_definition Rep Abs { α . [w0  ψ{α}]}
      and AOT_rigid_restriction_condition ψ
  shows AOT_restricted_type ψ Rep Abs
  by (auto intro!: assms AOT_restricted_type_axioms.intro AOT_restricted_type.intro)



MLfun register_rigid_restricted_type (name:string, restriction:string) thy =
let
val ctxt = thy
val ctxt = setupStrictWorld ctxt
val trm = Syntax.check_term ctxt (AOT_read_term @{nonterminal φ'} ctxt restriction)
val free = case (Term.add_frees trm []) of [f] => f
  | _ => raise Term.TERM ("Expected single free variable.", [trm])
val trm = Term.absfree free trm
val localeTerm = HOLogic.mk_Trueprop
  (Const (const_nameAOT_rigid_restriction_condition, dummyT) $ trm)
val localeTerm = Syntax.check_prop ctxt localeTerm
val int_bnd = Binding.concealed (Binding.qualify true "internal" (Binding.name name))
val bnds = {Rep_name = Binding.qualify true name (Binding.name "Rep"),
            Abs_name = Binding.qualify true "Abs" int_bnd,
            type_definition_name = Binding.qualify true "type_definition" int_bnd}

fun after_qed witts thy = let
  val thms = (map (Element.conclude_witness ctxt) (flat witts))
  
  val typeset = HOLogic.mk_Collect ("α", dummyT,
    constAOT_model_valid_in $ constw0 $
    (trm $ (Const (const_nameAOT_term_of_var, dummyT) $ Bound 0)))
  val typeset = Syntax.check_term thy typeset
  val nonempty_thm = Drule.OF
    (@{thm AOT_rigid_restriction_condition.type_set_nonempty}, thms)

  val ((_,st),thy) = Typedef.add_typedef {overloaded=true}
    (Binding.name name, [], Mixfix.NoSyn) typeset (SOME bnds)
    (fn ctxt => (Tactic.resolve_tac ctxt ([nonempty_thm]) 1)) thy
  val ({rep_type = _, abs_type = _, Rep_name = Rep_name, Abs_name = Abs_name,
        axiom_name = _},
     {inhabited = _, type_definition = type_definition, Rep = _,
      Rep_inverse = _, Abs_inverse = _, Rep_inject = _, Abs_inject = _,
      Rep_cases = _, Abs_cases = _, Rep_induct = _, Abs_induct = _}) = st

  val locale_thm = Drule.OF (@{thm AOT_restricted_type_intro}, type_definition::thms)

  val st = Interpretation.global_interpretation (([(@{locale AOT_restricted_type},
    ((name, true), (Expression.Named [
      ("ψ", trm),
      ("Rep", Const (Rep_name, dummyT)),
      ("Abs", Const (Abs_name, dummyT))], [])))
    ], [])) [] thy

  val st = Proof.refine_insert [locale_thm] st
  val thy = Proof.global_immediate_proof st

  val thy = Local_Theory.background_theory (AOT_Constraints.map (
    Symtab.update (name, (term_of (snd free), term_of (snd free))))) thy
  val thy = Local_Theory.background_theory (AOT_Restriction.map (
    Symtab.update (name, (trm, Const (Rep_name, dummyT))))) thy
  
  in thy end
in
Element.witness_proof after_qed [[localeTerm]] thy
end

val _ =
  Outer_Syntax.command
    command_keywordAOT_register_rigid_restricted_type
    "Register a restricted type."
    (((Parse.short_ident --| Parse.$$$ ":") -- Parse.term) >>
    (Toplevel.local_theory_to_proof NONE NONE o register_rigid_restricted_type));

(* Generalized mechanism for "AOT_restricted_type.∀I" followed by ∀E *)
MLfun get_instantiated_allI ctxt varname thm = let
val trm = Thm.concl_of thm
val trm = case trm of (@{const Trueprop} $ (@{const AOT_model_valid_in} $ _ $ x)) => x
                      | _ => raise Term.TERM ("Expected simple theorem.", [trm])
fun extractVars (Const (const_nameAOT_term_of_var, t) $ (Const rep $ Var v)) =
    (if fst (fst v) = fst varname
     then [Const (const_nameAOT_term_of_var, t) $ (Const rep $ Var v)]
     else []) (* TODO: care about the index *)
  | extractVars (t1 $ t2) = extractVars t1 @ extractVars t2
  | extractVars (Abs (_, _, t)) = extractVars t
  | extractVars _ = []
val vars = extractVars trm
val vartrm = hd vars
val vars = fold Term.add_vars vars []
val var = hd vars
val trmty = (case vartrm of (Const (_, Type ("fun", [_, ty])) $ _) => ty
                       | _ => raise Match)
val varty = snd var
val tyname = fst (Term.dest_Type varty)
val b = tyname^".∀I" (* TODO: better way to find the theorem *)
val thms = fst (Context.map_proof_result (fn ctxt => (Attrib.eval_thms ctxt
    [(Facts.Named ((b,Position.none),NONE),[])], ctxt)) ctxt)
val allthm = (case thms of (thm::_) => thm
    | _ => raise Fail "Unknown restricted type.")
val trm = Abs (Term.string_of_vname (fst var), trmty, Term.abstract_over (vartrm, trm))
val trm = Thm.cterm_of (Context.proof_of ctxt) trm
val phi = hd (Term.add_vars (Thm.prop_of allthm) [])
val allthm = Drule.instantiate_normalize (TVars.empty, Vars.make [(phi,trm)]) allthm
in
allthm
end

(* TODO: unconstraining multiple variables does not work yet *)
attribute_setup "unconstrain" =
  Scan.lift (Scan.repeat1 Args.var) >> (fn args => Thm.rule_attribute []
  (fn ctxt => fn thm =>
    let
    val thm = fold (fn arg => fn thm => thm RS get_instantiated_allI ctxt arg thm)
                   args thm
    val thm = fold (fn _ => fn thm => thm RS @{thm "∀E"(2)}) args thm
    in
     thm
    end))
  "Generalize a statement about restricted variables to a statement about
   unrestricted variables with explicit restriction condition."



context AOT_restricted_type
begin

AOT_theorem "rule-ui":
  assumes α(ψ{α}  φ{α})
  shows φ{«AOT_term_of_var (Rep α)»}
proof -
  AOT_have φ{α} if ψ{α} for α using assms[THEN "∀E"(2), THEN "→E"] that by blast
  moreover AOT_have ψ{«AOT_term_of_var (Rep α)»}
    by (auto simp:  ψ)
  ultimately show ?thesis by blast
qed
lemmas "∀E" = "rule-ui"

AOT_theorem "instantiation":
  assumes for arbitrary β: φ{«AOT_term_of_var (Rep β)»}  χ and α (ψ{α} & φ{α})
  shows χ
proof -
  AOT_have φ{«AOT_term_of_var (Rep α)»}  χ for α
    using assms(1)
    by (simp add: "deduction-theorem")
  AOT_hence 0: α (ψ{α}  (φ{α}  χ))
    using GEN by simp
  moreover AOT_obtain α where ψ{α} & φ{α} using assms(2) "∃E"[rotated] by blast
  ultimately AOT_show χ using "AOT_PLM.∀E"(2)[THEN "→E", THEN "→E"] "&E" by fast
qed
lemmas "∃E" = "instantiation"

AOT_theorem existential: assumes φ{«AOT_term_of_var (Rep β)»}
  shows  α (ψ{α} & φ{α})
  by (meson AOT_restricted_type.ψ AOT_restricted_type_axioms assms
            "&I" "existential:2[const_var]")
lemmas "∃I" = existential
end


context AOT_rigid_restriction_condition
begin

AOT_theorem "res-var-bound-reas[1]":
  α(ψ{α}  β φ{α, β})  βα (ψ{α}  φ{α, β})
proof(safe intro!: "≡I" "→I" GEN)
  fix β α
  AOT_assume α (ψ{α}  β φ{α, β})
  AOT_hence ψ{α}  β φ{α, β} using "∀E"(2) by blast
  moreover AOT_assume ψ{α}
  ultimately AOT_have β φ{α, β} using "→E" by blast
  AOT_thus φ{α, β} using "∀E"(2) by blast
next
  fix α β
  AOT_assume βα(ψ{α}  φ{α, β})
  AOT_hence α(ψ{α}  φ{α, β}) using "∀E"(2) by blast
  AOT_hence ψ{α}  φ{α, β} using "∀E"(2) by blast
  moreover AOT_assume ψ{α}
  ultimately AOT_show φ{α, β} using "→E" by blast
qed

AOT_theorem "res-var-bound-reas[BF]":
  α(ψ{α}  φ{α})  α(ψ{α}  φ{α})
proof(safe intro!: "→I")
  AOT_assume α(ψ{α}  φ{α})
  AOT_hence ψ{α}  φ{α} for α
    using "∀E"(2) by blast
  AOT_hence (ψ{α}  φ{α}) for α
    by (metis "sc-eq-box-box:6" rigid_condition "vdash-properties:6")
  AOT_hence α (ψ{α}  φ{α})
    by (rule GEN)
  AOT_thus α (ψ{α}  φ{α})
    by (metis "BF" "vdash-properties:6")
qed

AOT_theorem "res-var-bound-reas[CBF]":
α(ψ{α}  φ{α})  α(ψ{α}  φ{α})
proof(safe intro!: "→I" GEN)
  fix α
  AOT_assume α (ψ{α}  φ{α})
  AOT_hence α (ψ{α}  φ{α})
    by (metis "CBF" "vdash-properties:6")
  AOT_hence 1: (ψ{α}  φ{α})
    using "∀E"(2) by blast
  AOT_assume ψ{α}
  AOT_hence ψ{α}
    by (metis "B◇" "T◇" rigid_condition "vdash-properties:6")
  AOT_thus φ{α}
    using 1 "qml:1"[axiom_inst, THEN "→E", THEN "→E"] by blast
qed

AOT_theorem "res-var-bound-reas[2]":
α (ψ{α}  𝒜φ{α})  𝒜α (ψ{α}  φ{α})
proof(safe intro!: "→I")
  AOT_assume α (ψ{α}  𝒜φ{α})
  AOT_hence ψ{α}  𝒜φ{α} for α
    using "∀E"(2) by blast
  AOT_hence 𝒜(ψ{α}  φ{α}) for α
    by (metis "sc-eq-box-box:7" rigid_condition "vdash-properties:6")
  AOT_hence α 𝒜(ψ{α}  φ{α})
    by (rule GEN)
  AOT_thus 𝒜α(ψ{α}  φ{α})
    by (metis "≡E"(2) "logic-actual-nec:3"[axiom_inst])
qed


AOT_theorem "res-var-bound-reas[3]":
  𝒜α (ψ{α}  φ{α})  α (ψ{α}  𝒜φ{α})
proof(safe intro!: "→I" GEN)
  fix α
  AOT_assume 𝒜α (ψ{α}  φ{α})
  AOT_hence α 𝒜(ψ{α}  φ{α})
    by (metis "≡E"(1) "logic-actual-nec:3"[axiom_inst])
  AOT_hence 1: 𝒜(ψ{α}  φ{α}) by (metis "rule-ui:3")
  AOT_assume ψ{α}
  AOT_hence 𝒜ψ{α}
    by (metis "nec-imp-act" "qml:2"[axiom_inst] rigid_condition "→E")
  AOT_thus 𝒜φ{α}
    using 1 by (metis "act-cond" "→E")
qed

AOT_theorem "res-var-bound-reas[Buridan]":
  α (ψ{α} & φ{α})  α (ψ{α} & φ{α})
proof (rule "→I")
  AOT_assume α (ψ{α} & φ{α})
  then AOT_obtain α where ψ{α} & φ{α}
    using "∃E"[rotated] by blast
  AOT_hence (ψ{α} & φ{α}) 
    by (metis "KBasic:11" "KBasic:3" "T◇" "&I" "&E"(1) "&E"(2)
              "≡E"(2) "reductio-aa:1" rigid_condition "vdash-properties:6")
  AOT_hence α (ψ{α} & φ{α})
    by (rule "∃I")
  AOT_thus α (ψ{α} & φ{α})
    by (rule Buridan[THEN "→E"])
qed

AOT_theorem "res-var-bound-reas[BF◇]":
  α (ψ{α} & φ{α})  α (ψ{α} & φ{α})
proof(rule "→I")
  AOT_assume α (ψ{α} & φ{α})
  AOT_hence α (ψ{α} & φ{α})
    using "BF◇"[THEN "→E"] by blast
  then AOT_obtain α where (ψ{α} & φ{α})
    using "∃E"[rotated] by blast
  AOT_hence ψ{α} and φ{α}
    using "KBasic2:3" "&E" "→E" by blast+
  moreover AOT_have ψ{α}
    using calculation rigid_condition by (metis "B◇" "K◇" "→E")
  ultimately AOT_have ψ{α} & φ{α}
    using "&I" by blast
  AOT_thus α (ψ{α} & φ{α})
    by (rule "∃I")
qed

AOT_theorem "res-var-bound-reas[CBF◇]":
  α (ψ{α} & φ{α})  α (ψ{α} & φ{α})
proof(rule "→I")
  AOT_assume α (ψ{α} & φ{α})
  then AOT_obtain α where ψ{α} & φ{α}
    using "∃E"[rotated] by blast
  AOT_hence ψ{α} and φ{α}
    using rigid_condition[THEN "qml:2"[axiom_inst, THEN "→E"], THEN "→E"] "&E" by blast+
  AOT_hence (ψ{α} & φ{α})
    by (metis "KBasic:16" "con-dis-taut:5" "→E")
  AOT_hence α (ψ{α} & φ{α})
    by (rule "∃I")
  AOT_thus α (ψ{α} & φ{α})
    using "CBF◇"[THEN "→E"] by fast
qed

AOT_theorem "res-var-bound-reas[A-Exists:1]":
  𝒜∃!α (ψ{α} & φ{α})  ∃!α (ψ{α} & 𝒜φ{α})
proof(safe intro!: "≡I" "→I")
  AOT_assume 𝒜∃!α (ψ{α} & φ{α})
  AOT_hence ∃!α 𝒜(ψ{α} & φ{α})
    using "A-Exists:1"[THEN "≡E"(1)] by blast
  AOT_hence ∃!α (𝒜ψ{α} & 𝒜φ{α})
    apply(AOT_subst 𝒜ψ{α} & 𝒜φ{α} 𝒜(ψ{α} & φ{α}) for: α)
     apply (meson "Act-Basic:2" "intro-elim:3:f" "oth-class-taut:3:a")
    by simp
  AOT_thus ∃!α (ψ{α} & 𝒜φ{α})
    apply (AOT_subst ψ{α} 𝒜ψ{α} for: α)
    using "Commutativity of ≡" "intro-elim:3:b" "sc-eq-fur:2"
          "→E" rigid_condition by blast
next
  AOT_assume ∃!α (ψ{α} & 𝒜φ{α})
  AOT_hence ∃!α (𝒜ψ{α} & 𝒜φ{α})
    apply (AOT_subst 𝒜ψ{α} ψ{α} for: α)
     apply (meson "sc-eq-fur:2" "→E" rigid_condition)
    by simp
  AOT_hence ∃!α 𝒜(ψ{α} & φ{α})
    apply (AOT_subst 𝒜(ψ{α} & φ{α}) 𝒜ψ{α} & 𝒜φ{α} for: α)
     using "Act-Basic:2" apply presburger
     by simp
  AOT_thus 𝒜∃!α (ψ{α} & φ{α})
     by (metis "A-Exists:1" "intro-elim:3:b")
qed

end

(*<*)
end
(*>*)